src/HOL/Decision_Procs/MIR.thy
author haftmann
Thu, 22 Nov 2018 10:06:31 +0000
changeset 69325 4b6ddc5989fc
parent 69313 b021008c5397
child 69597 ff784d5a5bfb
permissions -rw-r--r--
removed legacy input syntax
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
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 66124
diff changeset
     7
  "HOL-Library.Code_Target_Numeral" "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
69325
4b6ddc5989fc removed legacy input syntax
haftmann
parents: 69313
diff changeset
    10
section \<open>Prelude\<close>
4b6ddc5989fc removed legacy input syntax
haftmann
parents: 69313
diff changeset
    11
4b6ddc5989fc removed legacy input syntax
haftmann
parents: 69313
diff changeset
    12
abbreviation (input) UNION :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> 'b set"
4b6ddc5989fc removed legacy input syntax
haftmann
parents: 69313
diff changeset
    13
  where "UNION A f \<equiv> \<Union> (f ` A)" \<comment> \<open>legacy\<close>
4b6ddc5989fc removed legacy input syntax
haftmann
parents: 69313
diff changeset
    14
4b6ddc5989fc removed legacy input syntax
haftmann
parents: 69313
diff changeset
    15
61586
5197a2ecb658 isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
    16
section \<open>Quantifier elimination for \<open>\<real> (0, 1, +, floor, <)\<close>\<close>
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
    17
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
    18
declare of_int_floor_cancel [simp del]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    19
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    20
lemma myle:
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    21
  fixes a b :: "'a::{ordered_ab_group_add}"
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
    22
  shows "(a \<le> b) = (0 \<le> b - a)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    23
  by (metis add_0_left add_le_cancel_right diff_add_cancel)
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    24
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    25
lemma myless:
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    26
  fixes a b :: "'a::{ordered_ab_group_add}"
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
    27
  shows "(a < b) = (0 < b - a)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    28
  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
    29
324622260d29 Added 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
(* Periodicity of dvd *)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
    31
lemmas dvd_period = zdvd_period
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    32
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
    33
(* The Divisibility relation between reals *)
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    34
definition rdvd:: "real \<Rightarrow> real \<Rightarrow> bool" (infixl "rdvd" 50)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
    35
  where "x rdvd y \<longleftrightarrow> (\<exists>k::int. y = x * real_of_int 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
    36
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
    37
lemma int_rdvd_real:
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
    38
  "real_of_int (i::int) rdvd x = (i dvd \<lfloor>x\<rfloor> \<and> real_of_int \<lfloor>x\<rfloor> = 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
    39
proof
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
    40
  assume "?l"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
    41
  hence th: "\<exists> k. x=real_of_int (i*k)" by (simp add: rdvd_def)
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
    42
  hence th': "real_of_int \<lfloor>x\<rfloor> = x" by (auto simp del: of_int_mult)
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
    43
  with th have "\<exists> k. real_of_int \<lfloor>x\<rfloor> = real_of_int (i*k)" by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
    44
  hence "\<exists>k. \<lfloor>x\<rfloor> = i*k" by presburger
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
    45
  thus ?r using th' by (simp add: dvd_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
    46
next
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60533
diff changeset
    47
  assume "?r" hence "(i::int) dvd \<lfloor>x::real\<rfloor>" ..
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
    48
  hence "\<exists>k. real_of_int \<lfloor>x\<rfloor> = real_of_int (i*k)"
61649
268d88ec9087 Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents: 61610
diff changeset
    49
    by (metis (no_types) dvd_def)
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
    50
  thus ?l using \<open>?r\<close> 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
    51
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
    52
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
    53
lemma int_rdvd_iff: "(real_of_int (i::int) rdvd real_of_int t) = (i dvd t)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
    54
  by (auto simp add: rdvd_def dvd_def) (rule_tac x="k" in exI, simp only: of_int_mult[symmetric])
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
    55
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
    56
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
    57
lemma rdvd_abs1: "(\<bar>real_of_int d\<bar> rdvd t) = (real_of_int (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
    58
proof
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
    59
  assume d: "real_of_int d rdvd t"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
    60
  from d int_rdvd_real have d2: "d dvd \<lfloor>t\<rfloor>" and ti: "real_of_int \<lfloor>t\<rfloor> = t"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    61
    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
    62
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
    63
  from iffD2[OF abs_dvd_iff] d2 have "\<bar>d\<bar> dvd \<lfloor>t\<rfloor>" by blast
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
    64
  with ti int_rdvd_real[symmetric] have "real_of_int \<bar>d\<bar> rdvd t" by blast
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
    65
  thus "\<bar>real_of_int d\<bar> rdvd 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
    66
next
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
    67
  assume "\<bar>real_of_int d\<bar> rdvd t" hence "real_of_int \<bar>d\<bar> rdvd t" by simp
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
    68
  with int_rdvd_real[where i="\<bar>d\<bar>" and x="t"]
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
    69
  have d2: "\<bar>d\<bar> dvd \<lfloor>t\<rfloor>" and ti: "real_of_int \<lfloor>t\<rfloor> = t"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    70
    by auto
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
    71
  from iffD1[OF abs_dvd_iff] d2 have "d dvd \<lfloor>t\<rfloor>" by blast
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
    72
  with ti int_rdvd_real[symmetric] show "real_of_int d rdvd 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
    73
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
    74
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
    75
lemma rdvd_minus: "(real_of_int (d::int) rdvd t) = (real_of_int d 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
    76
  apply (auto simp add: rdvd_def)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
    77
  apply (rule_tac x="-k" in exI, 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
    78
  apply (rule_tac x="-k" in exI, simp)
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    79
  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
    80
324622260d29 Added 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
lemma rdvd_left_0_eq: "(0 rdvd t) = (t=0)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    82
  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
    83
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
    84
lemma rdvd_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
    85
  assumes knz: "k\<noteq>0"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
    86
  shows "(real_of_int (n::int) * real_of_int (k::int) rdvd x * real_of_int k) = (real_of_int n rdvd x)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    87
  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
    88
324622260d29 Added 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
  (****                            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
    91
  (*********************************************************************************)
324622260d29 Added 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
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
    93
datatype (plugins del: size) num = C int | Bound nat | CN nat int num
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
    94
  | Neg num | Add num num | Sub num num
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
    95
  | Mul int num | Floor num | CF int num num
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
    96
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
    97
instantiation num :: size
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
    98
begin
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
    99
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   100
primrec size_num :: "num \<Rightarrow> nat"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   101
where
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   102
  "size_num (C c) = 1"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   103
| "size_num (Bound n) = 1"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   104
| "size_num (Neg a) = 1 + size_num a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   105
| "size_num (Add a b) = 1 + size_num a + size_num b"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   106
| "size_num (Sub a b) = 3 + size_num a + size_num b"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   107
| "size_num (CN n c a) = 4 + size_num a "
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   108
| "size_num (CF c a b) = 4 + size_num a + size_num b"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   109
| "size_num (Mul c a) = 1 + size_num a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   110
| "size_num (Floor a) = 1 + size_num a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   111
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   112
instance ..
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   113
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   114
end
23264
324622260d29 Added 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
324622260d29 Added 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
  (* Semantics of numeral terms (num) *)
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   117
primrec Inum :: "real list \<Rightarrow> num \<Rightarrow> real"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   118
where
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
   119
  "Inum bs (C c) = (real_of_int 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
   120
| "Inum bs (Bound n) = bs!n"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
   121
| "Inum bs (CN n c a) = (real_of_int c) * (bs!n) + (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
   122
| "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
   123
| "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
   124
| "Inum bs (Sub a b) = Inum bs a - Inum bs b"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
   125
| "Inum bs (Mul c a) = (real_of_int c) * Inum bs a"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
   126
| "Inum bs (Floor a) = real_of_int \<lfloor>Inum bs a\<rfloor>"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
   127
| "Inum bs (CF c a b) = real_of_int c * real_of_int \<lfloor>Inum bs a\<rfloor> + Inum bs b"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
   128
definition "isint t bs \<equiv> real_of_int \<lfloor>Inum bs t\<rfloor> = Inum bs t"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
   129
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
   130
lemma isint_iff: "isint n bs = (real_of_int \<lfloor>Inum bs n\<rfloor> = Inum bs n)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   131
  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
   132
324622260d29 Added 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
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
   134
  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
   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_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
   137
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
   138
  let ?e = "Inum bs e"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
   139
  assume be: "isint e bs" hence efe:"real_of_int \<lfloor>?e\<rfloor> = ?e" by (simp add: isint_iff)
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
   140
  have "real_of_int \<lfloor>Inum bs (Mul c e)\<rfloor> = real_of_int \<lfloor>real_of_int (c * \<lfloor>?e\<rfloor>)\<rfloor>"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
   141
    using efe by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
   142
  also have "\<dots> = real_of_int (c* \<lfloor>?e\<rfloor>)" by (metis floor_of_int)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
   143
  also have "\<dots> = real_of_int c * ?e" using efe 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
   144
  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
   145
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
   146
324622260d29 Added 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
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
   148
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
   149
  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
   150
  assume ie: "isint e bs"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
   151
  hence th: "real_of_int \<lfloor>?I e\<rfloor> = ?I e" by (simp add: isint_def)
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
   152
  have "real_of_int \<lfloor>?I (Neg e)\<rfloor> = real_of_int \<lfloor>- (real_of_int \<lfloor>?I e\<rfloor>)\<rfloor>"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
   153
    by (simp add: th)
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
   154
  also have "\<dots> = - real_of_int \<lfloor>?I e\<rfloor>" 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
   155
  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
   156
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
   157
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   158
lemma isint_sub:
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   159
  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
   160
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
   161
  let ?I = "\<lambda> t. Inum bs t"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
   162
  from ie have th: "real_of_int \<lfloor>?I e\<rfloor> = ?I e" by (simp add: isint_def)
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
   163
  have "real_of_int \<lfloor>?I (Sub (C c) e)\<rfloor> = real_of_int \<lfloor>real_of_int (c - \<lfloor>?I e\<rfloor>)\<rfloor>"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
   164
    by (simp add: th)
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
   165
  also have "\<dots> = real_of_int (c - \<lfloor>?I e\<rfloor>)" 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
   166
  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
   167
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
   168
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   169
lemma isint_add:
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   170
  assumes ai: "isint a bs" and bi: "isint b bs"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   171
  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
   172
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
   173
  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
   174
  let ?b = "Inum bs b"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
   175
  from ai bi isint_iff have "real_of_int \<lfloor>?a + ?b\<rfloor> = real_of_int \<lfloor>real_of_int \<lfloor>?a\<rfloor> + real_of_int \<lfloor>?b\<rfloor>\<rfloor>"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   176
    by simp
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
   177
  also have "\<dots> = real_of_int \<lfloor>?a\<rfloor> + real_of_int \<lfloor>?b\<rfloor>" 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
   178
  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
   179
  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
   180
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
   181
324622260d29 Added 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
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
   183
  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
324622260d29 Added 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
    (* FORMULAE *)
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   187
datatype (plugins del: size) fm =
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   188
  T | F | Lt num | Le num | Gt num | Ge num | Eq num | NEq num |
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   189
  Dvd int num | NDvd int num |
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   190
  NOT fm | And fm fm |  Or fm fm | Imp fm fm | Iff fm fm | E fm | A fm
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   191
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   192
instantiation fm :: size
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   193
begin
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   194
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   195
primrec size_fm :: "fm \<Rightarrow> nat"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   196
where
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   197
  "size_fm (NOT p) = 1 + size_fm p"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   198
| "size_fm (And p q) = 1 + size_fm p + size_fm q"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   199
| "size_fm (Or p q) = 1 + size_fm p + size_fm q"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   200
| "size_fm (Imp p q) = 3 + size_fm p + size_fm q"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   201
| "size_fm (Iff p q) = 3 + 2 * (size_fm p + size_fm q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   202
| "size_fm (E p) = 1 + size_fm p"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   203
| "size_fm (A p) = 4 + size_fm p"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   204
| "size_fm (Dvd i t) = 2"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   205
| "size_fm (NDvd i t) = 2"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   206
| "size_fm T = 1"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   207
| "size_fm F = 1"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   208
| "size_fm (Lt _) = 1"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   209
| "size_fm (Le _) = 1"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   210
| "size_fm (Gt _) = 1"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   211
| "size_fm (Ge _) = 1"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   212
| "size_fm (Eq _) = 1"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   213
| "size_fm (NEq _) = 1"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   214
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   215
instance ..
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   216
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   217
end
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   218
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   219
lemma size_fm_pos [simp]: "size p > 0" for p :: fm
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   220
  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
   221
324622260d29 Added 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
  (* Semantics of formulae (fm) *)
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   223
primrec Ifm ::"real list \<Rightarrow> fm \<Rightarrow> bool"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   224
where
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   225
  "Ifm bs T \<longleftrightarrow> True"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   226
| "Ifm bs F \<longleftrightarrow> False"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   227
| "Ifm bs (Lt a) \<longleftrightarrow> Inum bs a < 0"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   228
| "Ifm bs (Gt a) \<longleftrightarrow> Inum bs a > 0"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   229
| "Ifm bs (Le a) \<longleftrightarrow> Inum bs a \<le> 0"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   230
| "Ifm bs (Ge a) \<longleftrightarrow> Inum bs a \<ge> 0"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   231
| "Ifm bs (Eq a) \<longleftrightarrow> Inum bs a = 0"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   232
| "Ifm bs (NEq a) \<longleftrightarrow> Inum bs a \<noteq> 0"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   233
| "Ifm bs (Dvd i b) \<longleftrightarrow> real_of_int i rdvd Inum bs b"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   234
| "Ifm bs (NDvd i b) \<longleftrightarrow> \<not> (real_of_int i rdvd Inum bs b)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   235
| "Ifm bs (NOT p) \<longleftrightarrow> \<not> (Ifm bs p)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   236
| "Ifm bs (And p q) \<longleftrightarrow> Ifm bs p \<and> Ifm bs q"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   237
| "Ifm bs (Or p q) \<longleftrightarrow> Ifm bs p \<or> Ifm bs q"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   238
| "Ifm bs (Imp p q) \<longleftrightarrow> (Ifm bs p \<longrightarrow> Ifm bs q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   239
| "Ifm bs (Iff p q) \<longleftrightarrow> (Ifm bs p \<longleftrightarrow> Ifm bs q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   240
| "Ifm bs (E p) \<longleftrightarrow> (\<exists>x. Ifm (x # bs) p)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   241
| "Ifm bs (A p) \<longleftrightarrow> (\<forall>x. Ifm (x # bs) p)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   242
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   243
fun prep :: "fm \<Rightarrow> fm"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   244
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
   245
  "prep (E T) = T"
66124
7f0088571576 replaced recdef by fun
haftmann
parents: 66123
diff changeset
   246
| "prep (E F) = F"
7f0088571576 replaced recdef by fun
haftmann
parents: 66123
diff changeset
   247
| "prep (E (Or p q)) = Or (prep (E p)) (prep (E q))"
7f0088571576 replaced recdef by fun
haftmann
parents: 66123
diff changeset
   248
| "prep (E (Imp p q)) = Or (prep (E (NOT p))) (prep (E q))"
7f0088571576 replaced recdef by fun
haftmann
parents: 66123
diff changeset
   249
| "prep (E (Iff p q)) = Or (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))"
7f0088571576 replaced recdef by fun
haftmann
parents: 66123
diff changeset
   250
| "prep (E (NOT (And p q))) = Or (prep (E (NOT p))) (prep (E(NOT q)))"
7f0088571576 replaced recdef by fun
haftmann
parents: 66123
diff changeset
   251
| "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))"
7f0088571576 replaced recdef by fun
haftmann
parents: 66123
diff changeset
   252
| "prep (E (NOT (Iff p q))) = Or (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))"
7f0088571576 replaced recdef by fun
haftmann
parents: 66123
diff changeset
   253
| "prep (E p) = E (prep p)"
7f0088571576 replaced recdef by fun
haftmann
parents: 66123
diff changeset
   254
| "prep (A (And p q)) = And (prep (A p)) (prep (A q))"
7f0088571576 replaced recdef by fun
haftmann
parents: 66123
diff changeset
   255
| "prep (A p) = prep (NOT (E (NOT p)))"
7f0088571576 replaced recdef by fun
haftmann
parents: 66123
diff changeset
   256
| "prep (NOT (NOT p)) = prep p"
7f0088571576 replaced recdef by fun
haftmann
parents: 66123
diff changeset
   257
| "prep (NOT (And p q)) = Or (prep (NOT p)) (prep (NOT q))"
7f0088571576 replaced recdef by fun
haftmann
parents: 66123
diff changeset
   258
| "prep (NOT (A p)) = prep (E (NOT p))"
7f0088571576 replaced recdef by fun
haftmann
parents: 66123
diff changeset
   259
| "prep (NOT (Or p q)) = And (prep (NOT p)) (prep (NOT q))"
7f0088571576 replaced recdef by fun
haftmann
parents: 66123
diff changeset
   260
| "prep (NOT (Imp p q)) = And (prep p) (prep (NOT q))"
7f0088571576 replaced recdef by fun
haftmann
parents: 66123
diff changeset
   261
| "prep (NOT (Iff p q)) = Or (prep (And p (NOT q))) (prep (And (NOT p) q))"
7f0088571576 replaced recdef by fun
haftmann
parents: 66123
diff changeset
   262
| "prep (NOT p) = NOT (prep p)"
7f0088571576 replaced recdef by fun
haftmann
parents: 66123
diff changeset
   263
| "prep (Or p q) = Or (prep p) (prep q)"
7f0088571576 replaced recdef by fun
haftmann
parents: 66123
diff changeset
   264
| "prep (And p q) = And (prep p) (prep q)"
7f0088571576 replaced recdef by fun
haftmann
parents: 66123
diff changeset
   265
| "prep (Imp p q) = prep (Or (NOT p) q)"
7f0088571576 replaced recdef by fun
haftmann
parents: 66123
diff changeset
   266
| "prep (Iff p q) = Or (prep (And p q)) (prep (And (NOT p) (NOT q)))"
7f0088571576 replaced recdef by fun
haftmann
parents: 66123
diff changeset
   267
| "prep p = p"
7f0088571576 replaced recdef by fun
haftmann
parents: 66123
diff changeset
   268
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   269
lemma prep: "\<And> bs. Ifm bs (prep p) = Ifm bs p"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   270
  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
   271
324622260d29 Added 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
  (* Quantifier freeness *)
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   274
fun qfree:: "fm \<Rightarrow> bool"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   275
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
   276
  "qfree (E p) = False"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   277
| "qfree (A p) = False"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   278
| "qfree (NOT p) = qfree p"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   279
| "qfree (And p q) = (qfree p \<and> qfree q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   280
| "qfree (Or  p q) = (qfree p \<and> qfree q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   281
| "qfree (Imp p q) = (qfree p \<and> qfree q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   282
| "qfree (Iff p q) = (qfree p \<and> qfree q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   283
| "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
   284
324622260d29 Added 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
  (* Boundedness and substitution *)
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   286
primrec numbound0 :: "num \<Rightarrow> bool" (* a num is INDEPENDENT of Bound 0 *)
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   287
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
   288
  "numbound0 (C c) = True"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   289
| "numbound0 (Bound n) = (n>0)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   290
| "numbound0 (CN n i a) = (n > 0 \<and> numbound0 a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   291
| "numbound0 (Neg a) = numbound0 a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   292
| "numbound0 (Add a b) = (numbound0 a \<and> numbound0 b)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   293
| "numbound0 (Sub a b) = (numbound0 a \<and> numbound0 b)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   294
| "numbound0 (Mul i a) = numbound0 a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   295
| "numbound0 (Floor a) = numbound0 a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   296
| "numbound0 (CF c a b) = (numbound0 a \<and> numbound0 b)"
25765
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   297
23264
324622260d29 Added 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
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
   299
  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
   300
  shows "Inum (b#bs) a = Inum (b'#bs) a"
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
   301
  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
   302
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   303
lemma numbound0_gen:
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   304
  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
   305
  shows "\<forall> y. isint t (y#bs)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   306
  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
   307
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
   308
  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
   309
  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
   310
  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
   311
    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
   312
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
   313
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   314
primrec bound0:: "fm \<Rightarrow> bool" (* A Formula is independent of Bound 0 *)
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   315
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
   316
  "bound0 T = True"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   317
| "bound0 F = True"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   318
| "bound0 (Lt a) = numbound0 a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   319
| "bound0 (Le a) = numbound0 a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   320
| "bound0 (Gt a) = numbound0 a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   321
| "bound0 (Ge a) = numbound0 a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   322
| "bound0 (Eq a) = numbound0 a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   323
| "bound0 (NEq a) = numbound0 a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   324
| "bound0 (Dvd i a) = numbound0 a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   325
| "bound0 (NDvd i a) = numbound0 a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   326
| "bound0 (NOT p) = bound0 p"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   327
| "bound0 (And p q) = (bound0 p \<and> bound0 q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   328
| "bound0 (Or p q) = (bound0 p \<and> bound0 q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   329
| "bound0 (Imp p q) = ((bound0 p) \<and> (bound0 q))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   330
| "bound0 (Iff p q) = (bound0 p \<and> bound0 q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   331
| "bound0 (E p) = False"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   332
| "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
   333
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   334
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
   335
  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
   336
  shows "Ifm (b#bs) p = Ifm (b'#bs) p"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   337
  using bp numbound0_I [where b="b" and bs="bs" and b'="b'"]
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
   338
  by (induct p) auto
25765
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   339
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   340
primrec numsubst0:: "num \<Rightarrow> num \<Rightarrow> num" (* substitute a num into a num for Bound 0 *)
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   341
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
  "numsubst0 t (C c) = (C c)"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   343
| "numsubst0 t (Bound n) = (if n=0 then t else Bound n)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   344
| "numsubst0 t (CN n i a) = (if n=0 then Add (Mul i t) (numsubst0 t a) else CN n i (numsubst0 t a))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   345
| "numsubst0 t (CF i a b) = CF i (numsubst0 t a) (numsubst0 t b)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   346
| "numsubst0 t (Neg a) = Neg (numsubst0 t a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   347
| "numsubst0 t (Add a b) = Add (numsubst0 t a) (numsubst0 t b)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   348
| "numsubst0 t (Sub a b) = Sub (numsubst0 t a) (numsubst0 t b)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   349
| "numsubst0 t (Mul i a) = Mul i (numsubst0 t a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   350
| "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
   351
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   352
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
   353
  shows "Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b#bs) a)#bs) t"
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
   354
  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
   355
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   356
primrec subst0:: "num \<Rightarrow> fm \<Rightarrow> fm" (* substitue a num into a formula for Bound 0 *)
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   357
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
   358
  "subst0 t T = T"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   359
| "subst0 t F = F"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   360
| "subst0 t (Lt a) = Lt (numsubst0 t a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   361
| "subst0 t (Le a) = Le (numsubst0 t a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   362
| "subst0 t (Gt a) = Gt (numsubst0 t a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   363
| "subst0 t (Ge a) = Ge (numsubst0 t a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   364
| "subst0 t (Eq a) = Eq (numsubst0 t a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   365
| "subst0 t (NEq a) = NEq (numsubst0 t a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   366
| "subst0 t (Dvd i a) = Dvd i (numsubst0 t a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   367
| "subst0 t (NDvd i a) = NDvd i (numsubst0 t a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   368
| "subst0 t (NOT p) = NOT (subst0 t p)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   369
| "subst0 t (And p q) = And (subst0 t p) (subst0 t q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   370
| "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   371
| "subst0 t (Imp p q) = Imp (subst0 t p) (subst0 t q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   372
| "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
   373
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   374
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
   375
  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
   376
  using qfp numsubst0_I[where b="b" and bs="bs" and a="a"]
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
   377
  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
   378
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   379
fun decrnum:: "num \<Rightarrow> num"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   380
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
   381
  "decrnum (Bound n) = Bound (n - 1)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   382
| "decrnum (Neg a) = Neg (decrnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   383
| "decrnum (Add a b) = Add (decrnum a) (decrnum b)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   384
| "decrnum (Sub a b) = Sub (decrnum a) (decrnum b)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   385
| "decrnum (Mul c a) = Mul c (decrnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   386
| "decrnum (Floor a) = Floor (decrnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   387
| "decrnum (CN n c a) = CN (n - 1) c (decrnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   388
| "decrnum (CF c a b) = CF c (decrnum a) (decrnum b)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   389
| "decrnum a = a"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   390
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   391
fun decr :: "fm \<Rightarrow> fm"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   392
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
   393
  "decr (Lt a) = Lt (decrnum a)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   394
| "decr (Le a) = Le (decrnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   395
| "decr (Gt a) = Gt (decrnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   396
| "decr (Ge a) = Ge (decrnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   397
| "decr (Eq a) = Eq (decrnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   398
| "decr (NEq a) = NEq (decrnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   399
| "decr (Dvd i a) = Dvd i (decrnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   400
| "decr (NDvd i a) = NDvd i (decrnum a)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   401
| "decr (NOT p) = NOT (decr p)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   402
| "decr (And p q) = And (decr p) (decr q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   403
| "decr (Or p q) = Or (decr p) (decr q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   404
| "decr (Imp p q) = Imp (decr p) (decr q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   405
| "decr (Iff p q) = Iff (decr p) (decr q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   406
| "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
   407
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   408
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
   409
  shows "Inum (x#bs) t = Inum bs (decrnum t)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   410
  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
   411
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   412
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
   413
  shows "Ifm (x#bs) p = Ifm bs (decr p)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   414
  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
   415
324622260d29 Added 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
lemma decr_qf: "bound0 p \<Longrightarrow> qfree (decr p)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   417
  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
   418
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   419
fun isatom :: "fm \<Rightarrow> bool" (* test for atomicity *)
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   420
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
   421
  "isatom T = True"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   422
| "isatom F = True"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   423
| "isatom (Lt a) = True"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   424
| "isatom (Le a) = True"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   425
| "isatom (Gt a) = True"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   426
| "isatom (Ge a) = True"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   427
| "isatom (Eq a) = True"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   428
| "isatom (NEq a) = True"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   429
| "isatom (Dvd i b) = True"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   430
| "isatom (NDvd i b) = True"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   431
| "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
   432
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   433
lemma numsubst0_numbound0:
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   434
  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
   435
  shows "numbound0 (numsubst0 t a)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   436
  using nb by (induct a) auto
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   437
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   438
lemma subst0_bound0:
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   439
  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
   440
  shows "bound0 (subst0 t p)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   441
  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
   442
324622260d29 Added 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
lemma bound0_qf: "bound0 p \<Longrightarrow> qfree p"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   444
  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
   445
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   446
25765
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   447
definition djf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm" where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   448
  "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
   449
  (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
   450
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   451
definition evaldjf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm" where
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   452
  "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
   453
324622260d29 Added 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
lemma djf_Or: "Ifm bs (djf f p q) = Ifm bs (Or (f p) q)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   455
  by (cases "q=T", simp add: djf_def,cases "q=F",simp add: djf_def)
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   456
  (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
   457
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   458
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
   459
  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
   460
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   461
lemma evaldjf_bound0:
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   462
  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
   463
  shows "bound0 (evaldjf f xs)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   464
  using nb
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   465
  apply (induct xs)
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   466
  apply (auto simp add: evaldjf_def djf_def Let_def)
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   467
  apply (case_tac "f a")
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   468
  apply auto
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   469
  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
   470
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   471
lemma evaldjf_qf:
23264
324622260d29 Added 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
  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
   473
  shows "qfree (evaldjf f xs)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   474
  using nb
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   475
  apply (induct xs)
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   476
  apply (auto simp add: evaldjf_def djf_def Let_def)
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   477
  apply (case_tac "f a")
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   478
  apply auto
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   479
  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
   480
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   481
fun disjuncts :: "fm \<Rightarrow> fm list"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   482
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
   483
  "disjuncts (Or p q) = (disjuncts p) @ (disjuncts q)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   484
| "disjuncts F = []"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   485
| "disjuncts p = [p]"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   486
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   487
fun conjuncts :: "fm \<Rightarrow> fm list"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   488
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
   489
  "conjuncts (And p q) = (conjuncts p) @ (conjuncts q)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   490
| "conjuncts T = []"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   491
| "conjuncts p = [p]"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   492
23264
324622260d29 Added 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
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
   494
  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
   495
324622260d29 Added 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
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
   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
  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
   499
  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
   500
    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
  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
   502
qed
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   503
23264
324622260d29 Added 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 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
   505
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
   506
  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
   507
  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
   508
    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
   509
  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
   510
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
   511
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
   512
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
   513
  "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
   514
324622260d29 Added 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
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
   516
  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
   517
  shows "Ifm bs (DJ f p) = Ifm bs (f p)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   518
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
   519
  have "Ifm bs (DJ f p) = (\<exists> q \<in> set (disjuncts p). Ifm bs (f q))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   520
    by (simp add: DJ_def evaldjf_ex)
23264
324622260d29 Added 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
  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
   522
  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
   523
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
   524
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   525
lemma DJ_qf: assumes
23264
324622260d29 Added 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
  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
   527
  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
   528
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
   529
  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
   530
  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
   531
  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
   532
  with fqf have th':"\<forall> q\<in> set (disjuncts p). qfree (f q)" by blast
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   533
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   534
  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
   535
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
   536
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   537
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
   538
  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
   539
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
   540
  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
   541
  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
   542
  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
   543
  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
   544
  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
   545
    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
   546
  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
   547
  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
   548
  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
   549
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
   550
  (* 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
   551
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   552
  (* Algebraic simplifications for nums *)
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   553
fun bnds:: "num \<Rightarrow> nat list"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   554
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
   555
  "bnds (Bound n) = [n]"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   556
| "bnds (CN n c a) = n#(bnds a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   557
| "bnds (Neg a) = bnds a"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   558
| "bnds (Add a b) = (bnds a)@(bnds b)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   559
| "bnds (Sub a b) = (bnds a)@(bnds b)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   560
| "bnds (Mul i a) = bnds a"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   561
| "bnds (Floor a) = bnds a"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   562
| "bnds (CF c a b) = (bnds a)@(bnds b)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   563
| "bnds a = []"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   564
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   565
fun lex_ns:: "nat list \<Rightarrow> nat list \<Rightarrow> bool"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   566
where
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   567
  "lex_ns [] ms = True"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   568
| "lex_ns ns [] = False"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   569
| "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
   570
definition lex_bnd :: "num \<Rightarrow> num \<Rightarrow> bool" where
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   571
  "lex_bnd t s \<equiv> lex_ns (bnds t) (bnds s)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   572
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   573
fun maxcoeff:: "num \<Rightarrow> int"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   574
where
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   575
  "maxcoeff (C i) = \<bar>i\<bar>"
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   576
| "maxcoeff (CN n c t) = max \<bar>c\<bar> (maxcoeff t)"
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   577
| "maxcoeff (CF c t s) = max \<bar>c\<bar> (maxcoeff s)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   578
| "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
   579
324622260d29 Added 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
lemma maxcoeff_pos: "maxcoeff t \<ge> 0"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   581
  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
   582
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   583
fun numgcdh:: "num \<Rightarrow> int \<Rightarrow> int"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   584
where
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   585
  "numgcdh (C i) = (\<lambda>g. gcd i g)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   586
| "numgcdh (CN n c t) = (\<lambda>g. gcd c (numgcdh t g))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   587
| "numgcdh (CF c s t) = (\<lambda>g. gcd c (numgcdh t g))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   588
| "numgcdh t = (\<lambda>g. 1)"
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
   589
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   590
definition numgcd :: "num \<Rightarrow> int"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   591
  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
   592
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   593
fun reducecoeffh:: "num \<Rightarrow> int \<Rightarrow> num"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   594
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
   595
  "reducecoeffh (C i) = (\<lambda> g. C (i div g))"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   596
| "reducecoeffh (CN n c t) = (\<lambda> g. CN n (c div g) (reducecoeffh t g))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   597
| "reducecoeffh (CF c s t) = (\<lambda> g. CF (c div g)  s (reducecoeffh t g))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   598
| "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
   599
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   600
definition reducecoeff :: "num \<Rightarrow> num"
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
   601
where
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   602
  "reducecoeff t =
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   603
    (let g = numgcd t in
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   604
     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
   605
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   606
fun dvdnumcoeff:: "num \<Rightarrow> int \<Rightarrow> bool"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   607
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
   608
  "dvdnumcoeff (C i) = (\<lambda> g. g dvd i)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   609
| "dvdnumcoeff (CN n c t) = (\<lambda> g. g dvd c \<and> (dvdnumcoeff t g))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   610
| "dvdnumcoeff (CF c s t) = (\<lambda> g. g dvd c \<and> (dvdnumcoeff t g))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   611
| "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
   612
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   613
lemma dvdnumcoeff_trans:
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   614
  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
   615
  shows "dvdnumcoeff t g"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   616
  using dgt' gdg
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   617
  by (induct t rule: dvdnumcoeff.induct) (simp_all add: gdg dvd_trans[OF gdg])
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
   618
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
   619
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
   620
324622260d29 Added 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
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
   622
  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
   623
  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
   624
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
   625
  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
   626
    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
   627
  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
   628
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
   629
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   630
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
   631
  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
   632
324622260d29 Added 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
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
   634
  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
   635
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   636
lemma reducecoeffh:
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   637
  assumes gt: "dvdnumcoeff t g" and gp: "g > 0"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
   638
  shows "real_of_int g *(Inum bs (reducecoeffh t g)) = Inum bs 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
   639
  using gt
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   640
proof(induct t rule: reducecoeffh.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
   641
  case (1 i) hence gd: "g dvd i" by simp
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
   642
  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
   643
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
   644
  case (2 n c t)  hence gd: "g dvd c" by simp
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
   645
  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
   646
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
   647
  case (3 c s t)  hence gd: "g dvd c" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   648
  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
   649
qed (auto simp add: numgcd_def gp)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   650
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   651
fun ismaxcoeff:: "num \<Rightarrow> int \<Rightarrow> bool"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   652
where
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   653
  "ismaxcoeff (C i) = (\<lambda> x. \<bar>i\<bar> \<le> x)"
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   654
| "ismaxcoeff (CN n c t) = (\<lambda>x. \<bar>c\<bar> \<le> x \<and> (ismaxcoeff t x))"
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   655
| "ismaxcoeff (CF c s t) = (\<lambda>x. \<bar>c\<bar> \<le> x \<and> (ismaxcoeff t x))"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   656
| "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
   657
324622260d29 Added 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
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
   659
  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
   660
324622260d29 Added 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
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
   662
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
   663
  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
   664
  hence H:"ismaxcoeff t (maxcoeff t)" .
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   665
  have thh: "maxcoeff t \<le> max \<bar>c\<bar> (maxcoeff t)" by simp
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   666
  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
   667
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   668
  case (3 c t 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
   669
  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
   670
  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
   671
  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
   672
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
   673
67118
ccab07d1196c more simplification rules
haftmann
parents: 66809
diff changeset
   674
lemma zgcd_gt1:
ccab07d1196c more simplification rules
haftmann
parents: 66809
diff changeset
   675
  "\<bar>i\<bar> > 1 \<and> \<bar>j\<bar> > 1 \<or> \<bar>i\<bar> = 0 \<and> \<bar>j\<bar> > 1 \<or> \<bar>i\<bar> > 1 \<and> \<bar>j\<bar> = 0"
ccab07d1196c more simplification rules
haftmann
parents: 66809
diff changeset
   676
  if "gcd i j > 1" for i j :: int
ccab07d1196c more simplification rules
haftmann
parents: 66809
diff changeset
   677
proof -
ccab07d1196c more simplification rules
haftmann
parents: 66809
diff changeset
   678
  have "\<bar>k\<bar> \<le> 1 \<longleftrightarrow> k = - 1 \<or> k = 0 \<or> k = 1" for k :: int
ccab07d1196c more simplification rules
haftmann
parents: 66809
diff changeset
   679
    by auto
ccab07d1196c more simplification rules
haftmann
parents: 66809
diff changeset
   680
  with that show ?thesis
ccab07d1196c more simplification rules
haftmann
parents: 66809
diff changeset
   681
    by (auto simp add: not_less)
ccab07d1196c more simplification rules
haftmann
parents: 66809
diff changeset
   682
qed
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   683
23264
324622260d29 Added 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
lemma numgcdh0:"numgcdh t m = 0 \<Longrightarrow>  m =0"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   685
  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
   686
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   687
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
   688
  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
   689
  shows "dvdnumcoeff t (numgcdh t m)"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   690
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
   691
proof(induct t rule: numgcdh.induct)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   692
  case (2 n 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
   693
  let ?g = "numgcdh t m"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   694
  from 2 have th:"gcd c ?g > 1" by simp
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27456
diff changeset
   695
  from zgcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"]
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   696
  have "(\<bar>c\<bar> > 1 \<and> ?g > 1) \<or> (\<bar>c\<bar> = 0 \<and> ?g > 1) \<or> (\<bar>c\<bar> > 1 \<and> ?g = 0)" by simp
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   697
  moreover {assume "\<bar>c\<bar> > 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
   698
    have th: "dvdnumcoeff t ?g" by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   699
    have th': "gcd c ?g dvd ?g" by simp
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   700
    from dvdnumcoeff_trans[OF th' th] have ?case by simp }
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   701
  moreover {assume "\<bar>c\<bar> = 0 \<and> ?g > 1"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   702
    with 2 have th: "dvdnumcoeff t ?g" by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   703
    have th': "gcd c ?g dvd ?g" by simp
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   704
    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
   705
    hence ?case by simp }
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   706
  moreover {assume "\<bar>c\<bar> > 1" and g0:"?g = 0"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   707
    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
   708
  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
   709
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   710
  case (3 c s 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
   711
  let ?g = "numgcdh t m"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   712
  from 3 have th:"gcd c ?g > 1" by simp
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27456
diff changeset
   713
  from zgcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"]
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   714
  have "(\<bar>c\<bar> > 1 \<and> ?g > 1) \<or> (\<bar>c\<bar> = 0 \<and> ?g > 1) \<or> (\<bar>c\<bar> > 1 \<and> ?g = 0)" by simp
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   715
  moreover {assume "\<bar>c\<bar> > 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
   716
    have th: "dvdnumcoeff t ?g" by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   717
    have th': "gcd c ?g dvd ?g" by simp
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   718
    from dvdnumcoeff_trans[OF th' th] have ?case by simp }
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   719
  moreover {assume "\<bar>c\<bar> = 0 \<and> ?g > 1"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   720
    with 3 have th: "dvdnumcoeff t ?g" by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   721
    have th': "gcd c ?g dvd ?g" by simp
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   722
    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
   723
    hence ?case by simp }
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   724
  moreover {assume "\<bar>c\<bar> > 1" and g0:"?g = 0"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   725
    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
   726
  ultimately show ?case by blast
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   727
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
   728
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   729
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
   730
  assumes "numgcd t > 1" shows "dvdnumcoeff t (numgcd t) \<and> numgcd t > 0"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   731
  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
   732
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
   733
  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
   734
  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
   735
  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
   736
  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
   737
  assume H: "numgcdh t ?mc > 1"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   738
  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
   739
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
   740
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
   741
lemma reducecoeff: "real_of_int (numgcd t) * (Inum bs (reducecoeff t)) = Inum bs 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
   742
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
   743
  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
   744
  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
   745
  hence "?g = 0 \<or> ?g = 1 \<or> ?g > 1" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   746
  moreover {assume "?g = 0" hence ?thesis by (simp add: numgcd0)}
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   747
  moreover {assume "?g = 1" hence ?thesis by (simp add: reducecoeff_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
   748
  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
   749
    from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff t ?g" and g0: "?g > 0" by blast+
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   750
    from reducecoeffh[OF th1 g0, where bs="bs"] g1 have ?thesis
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   751
      by (simp add: reducecoeff_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
   752
  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
   753
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
   754
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   755
lemma reducecoeffh_numbound0: "numbound0 t \<Longrightarrow> numbound0 (reducecoeffh t g)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   756
  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
   757
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   758
lemma reducecoeff_numbound0: "numbound0 t \<Longrightarrow> numbound0 (reducecoeff t)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   759
  using reducecoeffh_numbound0 by (simp add: reducecoeff_def Let_def)
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   760
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   761
consts numadd:: "num \<times> num \<Rightarrow> num"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   762
recdef numadd "measure (\<lambda>(t, s). size t + size 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
   763
  "numadd (CN n1 c1 r1,CN n2 c2 r2) =
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   764
  (if n1=n2 then
23264
324622260d29 Added 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
  (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
   766
  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
   767
  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
   768
  else (CN n2 c2 (numadd (CN n1 c1 r1,r2))))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   769
  "numadd (CN n1 c1 r1,t) = CN n1 c1 (numadd (r1, t))"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   770
  "numadd (t,CN n2 c2 r2) = CN n2 c2 (numadd (t,r2))"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   771
  "numadd (CF c1 t1 r1,CF c2 t2 r2) =
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   772
   (if t1 = t2 then
23264
324622260d29 Added 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
    (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
   774
   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
   775
   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
   776
  "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
   777
  "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
   778
  "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
   779
  "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
   780
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   781
lemma numadd [simp]: "Inum bs (numadd (t, s)) = Inum bs (Add t s)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   782
  by (induct t s rule: numadd.induct) (simp_all add: Let_def algebra_simps add_eq_0_iff)
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   783
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   784
lemma numadd_nb [simp]: "numbound0 t \<Longrightarrow> numbound0 s \<Longrightarrow> numbound0 (numadd (t, s))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   785
  by (induct t s rule: numadd.induct) (simp_all add: Let_def)
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   786
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   787
fun nummul:: "num \<Rightarrow> int \<Rightarrow> num"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   788
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
   789
  "nummul (C j) = (\<lambda> i. C (i*j))"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   790
| "nummul (CN n c t) = (\<lambda> i. CN n (c*i) (nummul t i))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   791
| "nummul (CF c t s) = (\<lambda> i. CF (c*i) t (nummul s i))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   792
| "nummul (Mul c t) = (\<lambda> i. nummul t (i*c))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   793
| "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
   794
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   795
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
   796
  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
   797
324622260d29 Added 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
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
   799
  by (induct t rule: nummul.induct) auto
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   800
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   801
definition numneg :: "num \<Rightarrow> num"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   802
  where "numneg t \<equiv> nummul t (- 1)"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   803
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   804
definition numsub :: "num \<Rightarrow> num \<Rightarrow> num"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   805
  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
   806
324622260d29 Added 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
lemma numneg[simp]: "Inum bs (numneg t) = Inum bs (Neg t)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   808
  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
   809
324622260d29 Added 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
lemma numneg_nb[simp]: "numbound0 t \<Longrightarrow> numbound0 (numneg t)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   811
  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
   812
324622260d29 Added 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
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
   814
  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
   815
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   816
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
   817
  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
   818
324622260d29 Added 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
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
   820
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
   821
  have cti: "isint (Mul c (Floor t)) bs" by (simp add: isint_Mul isint_Floor)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   822
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   823
  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
   824
  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
   825
  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
   826
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
   827
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   828
fun split_int:: "num \<Rightarrow> num \<times> num"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   829
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
   830
  "split_int (C c) = (C 0, C c)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   831
| "split_int (CN n c b) =
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   832
     (let (bv,bi) = split_int 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
   833
       in (CN n c bv, bi))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   834
| "split_int (CF c a b) =
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   835
     (let (bv,bi) = split_int 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
   836
       in (bv, CF c a bi))"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   837
| "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
   838
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   839
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
   840
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
   841
  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
   842
  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
   843
  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
   844
  have "split_int b = (?bv,?bi)" by simp
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   845
  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
   846
  from 2(2) have tibi: "ti = ?bi" by (simp add: Let_def split_def)
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   847
  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
   848
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   849
  case (3 c a b tv 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
   850
  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
   851
  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
   852
  have "split_int b = (?bv,?bi)" by simp
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   853
  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
   854
  from 3(2) have tibi: "ti = CF c a ?bi"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   855
    by (simp add: Let_def split_def)
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   856
  from 3(2) b[symmetric] bii show ?case
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   857
    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
   858
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
   859
324622260d29 Added 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
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
   861
  by (induct t rule: split_int.induct) (auto simp add: Let_def split_def)
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   862
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   863
definition numfloor:: "num \<Rightarrow> num"
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
   864
where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   865
  "numfloor t = (let (tv,ti) = split_int t in
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   866
  (case tv of C i \<Rightarrow> numadd (tv,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
   867
  | _ \<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
   868
324622260d29 Added 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
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
   870
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
   871
  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
   872
  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
   873
  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
   874
  {assume H: "\<forall> v. ?tv \<noteq> C v"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   875
    hence th1: "?n t = ?N (Add (Floor ?tv) ?ti)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   876
      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
   877
    from split_int[OF tvti] have "?N (Floor t) = ?N (Floor(Add ?tv ?ti))" and tii:"isint ?ti bs" by simp+
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
   878
    hence "?N (Floor t) = real_of_int \<lfloor>?N (Add ?tv ?ti)\<rfloor>" by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
   879
    also have "\<dots> = real_of_int (\<lfloor>?N ?tv\<rfloor> + \<lfloor>?N ?ti\<rfloor>)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   880
      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
   881
    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
   882
    finally have ?thesis using th1 by simp}
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   883
  moreover {fix v assume H:"?tv = C v"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   884
    from split_int[OF tvti] have "?N (Floor t) = ?N (Floor(Add ?tv ?ti))" and tii:"isint ?ti bs" by simp+
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
   885
    hence "?N (Floor t) = real_of_int \<lfloor>?N (Add ?tv ?ti)\<rfloor>" by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
   886
    also have "\<dots> = real_of_int (\<lfloor>?N ?tv\<rfloor> + \<lfloor>?N ?ti\<rfloor>)"
23264
324622260d29 Added 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
      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
   888
    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
   889
    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
   890
  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
   891
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
   892
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   893
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
   894
  using split_int_nb[where t="t"]
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   895
  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
   896
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   897
fun simpnum:: "num \<Rightarrow> num"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   898
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
   899
  "simpnum (C j) = C j"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   900
| "simpnum (Bound n) = CN n 1 (C 0)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   901
| "simpnum (Neg t) = numneg (simpnum t)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   902
| "simpnum (Add t s) = numadd (simpnum t,simpnum s)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   903
| "simpnum (Sub t s) = numsub (simpnum t) (simpnum s)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   904
| "simpnum (Mul i t) = (if i = 0 then (C 0) else nummul (simpnum t) i)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   905
| "simpnum (Floor t) = numfloor (simpnum t)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   906
| "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
   907
| "simpnum (CF c t s) = simpnum(Add (Mul c (Floor t)) 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
   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 simpnum_ci[simp]: "Inum bs (simpnum t) = Inum bs t"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   910
  by (induct t rule: simpnum.induct) auto
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   911
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   912
lemma simpnum_numbound0[simp]: "numbound0 t \<Longrightarrow> numbound0 (simpnum t)"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   913
  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
   914
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   915
fun nozerocoeff:: "num \<Rightarrow> bool"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
   916
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
   917
  "nozerocoeff (C c) = True"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   918
| "nozerocoeff (CN n c t) = (c\<noteq>0 \<and> nozerocoeff t)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   919
| "nozerocoeff (CF c s t) = (c \<noteq> 0 \<and> nozerocoeff t)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   920
| "nozerocoeff (Mul c t) = (c\<noteq>0 \<and> nozerocoeff t)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   921
| "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
   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 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
   924
  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
   925
324622260d29 Added 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
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
   927
  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
   928
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   929
lemma numneg_nz : "nozerocoeff a \<Longrightarrow> nozerocoeff (numneg a)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   930
  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
   931
324622260d29 Added 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
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
   933
  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
   934
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   935
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
   936
  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
   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 numfloor_nz: "nozerocoeff t \<Longrightarrow> nozerocoeff (numfloor t)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   939
  by (simp add: numfloor_def Let_def split_def)
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   940
    (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
   941
324622260d29 Added 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
lemma simpnum_nz: "nozerocoeff (simpnum t)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   943
  by (induct t rule: simpnum.induct)
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   944
    (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
   945
324622260d29 Added 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
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
   947
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
   948
  case (2 n c t)
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   949
  hence cnz: "c \<noteq>0" and mx: "max \<bar>c\<bar> (maxcoeff t) = 0" by simp+
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   950
  have "max \<bar>c\<bar> (maxcoeff t) \<ge> \<bar>c\<bar>" by simp
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   951
  with cnz have "max \<bar>c\<bar> (maxcoeff t) > 0" by arith
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   952
  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
   953
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   954
  case (3 c s t)
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   955
  hence cnz: "c \<noteq>0" and mx: "max \<bar>c\<bar> (maxcoeff t) = 0" by simp+
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   956
  have "max \<bar>c\<bar> (maxcoeff t) \<ge> \<bar>c\<bar>" by simp
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   957
  with cnz have "max \<bar>c\<bar> (maxcoeff t) > 0" by arith
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   958
  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
   959
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
   960
324622260d29 Added 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
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
   962
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
   963
  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
   964
  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
   965
  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
   966
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
   967
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
   968
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
   969
  "simp_num_pair \<equiv> (\<lambda> (t,n). (if n = 0 then (C 0, 0) else
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   970
   (let t' = simpnum t ; g = numgcd t' in
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   971
      if g > 1 then (let g' = gcd n g in
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   972
        if g' = 1 then (t',n)
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
   973
        else (reducecoeffh t' g', n div 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
   974
      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
   975
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   976
lemma simp_num_pair_ci:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
   977
  shows "((\<lambda> (t,n). Inum bs t / real_of_int n) (simp_num_pair (t,n))) = ((\<lambda> (t,n). Inum bs t / real_of_int 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
   978
  (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
   979
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
   980
  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
   981
  let ?g = "numgcd ?t'"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   982
  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
   983
  {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
   984
  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
   985
  { 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
   986
    {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
   987
    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
   988
    {assume g1:"?g>1" hence g0: "?g > 0" by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   989
      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
   990
      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
   991
      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
   992
      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
   993
      moreover {assume g'1:"?g'>1"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
   994
        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
   995
        let ?tt = "reducecoeffh ?t' ?g'"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
   996
        let ?t = "Inum bs ?tt"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
   997
        have gpdg: "?g' dvd ?g" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
   998
        have gpdd: "?g' dvd n" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
   999
        have gpdgp: "?g' dvd ?g'" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1000
        from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1001
        have th2:"real_of_int ?g' * ?t = Inum bs ?t'" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1002
        from nnz g1 g'1 have "?lhs = ?t / real_of_int (n div ?g')" by (simp add: simp_num_pair_def Let_def)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1003
        also have "\<dots> = (real_of_int ?g' * ?t) / (real_of_int ?g' * (real_of_int (n div ?g')))" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1004
        also have "\<dots> = (Inum bs ?t' / real_of_int n)"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  1005
          using real_of_int_div[OF gpdd] th2 gp0 by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1006
        finally have "?lhs = Inum bs t / real_of_int n" by simp
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1007
        then have ?thesis using nnz g1 g'1 by (simp add: simp_num_pair_def) }
68270
2bc921b2159b treat gcd_eq_1_imp_coprime analogously to mod_0_imp_dvd
haftmann
parents: 67613
diff changeset
  1008
      ultimately have ?thesis by auto }
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1009
    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
  1010
  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
  1011
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
  1012
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1013
lemma simp_num_pair_l:
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1014
  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
  1015
  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
  1016
proof-
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1017
  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
  1018
  let ?g = "numgcd ?t'"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
  1019
  let ?g' = "gcd n ?g"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1020
  { 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
  1021
  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
  1022
  { assume nnz: "n \<noteq> 0"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1023
    {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
  1024
    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
  1025
    {assume g1:"?g>1" hence g0: "?g > 0" by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
  1026
      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
  1027
      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
  1028
      hence "?g'= 1 \<or> ?g' > 1" by arith
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1029
      moreover {assume "?g'=1" hence ?thesis using assms g1 g0
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1030
          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
  1031
      moreover {assume g'1:"?g'>1"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  1032
        have gpdg: "?g' dvd ?g" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  1033
        have gpdd: "?g' dvd n" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  1034
        have gpdgp: "?g' dvd ?g'" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  1035
        from zdvd_imp_le[OF gpdd np] have g'n: "?g' \<le> n" .
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  1036
        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
  1037
        have "n div ?g' >0" by simp
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1038
        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
  1039
          by(auto simp add: simp_num_pair_def Let_def reducecoeffh_numbound0)}
68270
2bc921b2159b treat gcd_eq_1_imp_coprime analogously to mod_0_imp_dvd
haftmann
parents: 67613
diff changeset
  1040
      ultimately have ?thesis by auto }
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1041
    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
  1042
  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
  1043
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
  1044
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1045
fun not:: "fm \<Rightarrow> fm"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1046
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
  1047
  "not (NOT p) = p"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1048
| "not T = F"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1049
| "not F = T"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1050
| "not (Lt t) = Ge t"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1051
| "not (Le t) = Gt t"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1052
| "not (Gt t) = Le t"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1053
| "not (Ge t) = Lt t"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1054
| "not (Eq t) = NEq t"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1055
| "not (NEq t) = Eq t"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1056
| "not (Dvd i t) = NDvd i t"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1057
| "not (NDvd i t) = Dvd i t"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1058
| "not (And p q) = Or (not p) (not q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1059
| "not (Or p q) = And (not p) (not q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1060
| "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
  1061
lemma not[simp]: "Ifm bs (not p) = Ifm bs (NOT p)"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1062
  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
  1063
lemma not_qf[simp]: "qfree p \<Longrightarrow> qfree (not p)"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1064
  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
  1065
lemma not_nb[simp]: "bound0 p \<Longrightarrow> bound0 (not p)"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1066
  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
  1067
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  1068
definition conj :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1069
  "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
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1070
   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
  1071
lemma conj[simp]: "Ifm bs (conj p q) = Ifm bs (And p q)"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1072
  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
  1073
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1074
lemma conj_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (conj p q)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1075
  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
  1076
lemma conj_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (conj p q)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1077
  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
  1078
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  1079
definition disj :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1080
  "disj p q \<equiv> (if (p = T \<or> q=T) then T else if p=F then q else if q=F then 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
  1081
       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
  1082
324622260d29 Added 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
lemma disj[simp]: "Ifm bs (disj p q) = Ifm bs (Or p q)"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1084
  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
  1085
lemma disj_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (disj p q)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1086
  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
  1087
lemma disj_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (disj p q)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1088
  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
  1089
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  1090
definition imp :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1091
  "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
23264
324622260d29 Added 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
    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
  1093
lemma imp[simp]: "Ifm bs (imp p q) = Ifm bs (Imp p q)"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1094
  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
  1095
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
  1096
  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
  1097
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  1098
definition iff :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1099
  "iff p q \<equiv> (if (p = q) then T else if (p = not q \<or> not p = q) then F else
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1100
       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
23264
324622260d29 Added 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
  Iff p q)"
61649
268d88ec9087 Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents: 61610
diff changeset
  1102
23264
324622260d29 Added 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
lemma iff[simp]: "Ifm bs (iff p q) = Ifm bs (Iff p q)"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1104
  by (unfold iff_def,cases "p=q", simp,cases "p=not q", simp)  (cases "not p= q", auto)
61649
268d88ec9087 Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents: 61610
diff changeset
  1105
23264
324622260d29 Added 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
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
  1107
  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
  1108
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1109
fun check_int:: "num \<Rightarrow> bool"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1110
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
  1111
  "check_int (C i) = True"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1112
| "check_int (Floor t) = True"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1113
| "check_int (Mul i t) = check_int t"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1114
| "check_int (Add t s) = (check_int t \<and> check_int s)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1115
| "check_int (Neg t) = check_int t"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1116
| "check_int (CF c t s) = check_int s"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1117
| "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
  1118
lemma check_int: "check_int t \<Longrightarrow> isint t bs"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  1119
  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
  1120
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1121
lemma rdvd_left1_int: "real_of_int \<lfloor>t\<rfloor> = t \<Longrightarrow> 1 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
  1122
  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
  1123
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1124
lemma rdvd_reduce:
23264
324622260d29 Added 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
  assumes gd:"g dvd d" and gc:"g dvd c" and gp: "g > 0"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1126
  shows "real_of_int (d::int) rdvd real_of_int (c::int)*t = (real_of_int (d div g) rdvd real_of_int (c div 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
  1127
proof
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1128
  assume d: "real_of_int d rdvd real_of_int c * t"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1129
  from d rdvd_def obtain k where k_def: "real_of_int c * t = real_of_int d* real_of_int (k::int)" 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
  1130
  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
  1131
  from gc dvd_def obtain kc where kc_def: "c = g * kc" by auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1132
  from k_def kd_def kc_def have "real_of_int g * real_of_int kc * t = real_of_int g * real_of_int kd * real_of_int k" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1133
  hence "real_of_int kc * t = real_of_int kd * real_of_int k" using gp by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1134
  hence th:"real_of_int kd rdvd real_of_int kc * t" using rdvd_def 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
  1135
  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
  1136
  from kc_def gp have "kc = c div g" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1137
  with th th' show "real_of_int (d div g) rdvd real_of_int (c div g) * 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
  1138
next
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1139
  assume d: "real_of_int (d div g) rdvd real_of_int (c div 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
  1140
  from gp have gnz: "g \<noteq> 0" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1141
  thus "real_of_int d rdvd real_of_int c * t" using d rdvd_mult[OF gnz, where n="d div g" and x="real_of_int (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
  1142
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
  1143
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  1144
definition simpdvd :: "int \<Rightarrow> num \<Rightarrow> (int \<times> num)" where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1145
  "simpdvd d t \<equiv>
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1146
   (let g = numgcd t in
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1147
      if g > 1 then (let g' = gcd d g in
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1148
        if g' = 1 then (d, t)
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1149
        else (d div g',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
  1150
      else (d, t))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1151
lemma simpdvd:
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1152
  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
  1153
  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
  1154
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
  1155
  let ?g = "numgcd t"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
  1156
  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
  1157
  {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
  1158
  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
  1159
  {assume g1:"?g>1" hence g0: "?g > 0" by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
  1160
    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
  1161
    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
  1162
    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
  1163
    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
  1164
    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
  1165
      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
  1166
      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
  1167
      let ?t = "Inum bs ?tt"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
  1168
      have gpdg: "?g' dvd ?g" by simp
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
  1169
      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
  1170
      have gpdgp: "?g' dvd ?g'" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1171
      from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1172
      have th2:"real_of_int ?g' * ?t = Inum bs t" by simp
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1173
      from assms g1 g0 g'1
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1174
      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
  1175
        by (simp add: simpdvd_def Let_def)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1176
      also have "\<dots> = (real_of_int d rdvd (Inum bs t))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1177
        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
  1178
          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
  1179
      finally have ?thesis by simp  }
68270
2bc921b2159b treat gcd_eq_1_imp_coprime analogously to mod_0_imp_dvd
haftmann
parents: 67613
diff changeset
  1180
    ultimately have ?thesis 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
  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
  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
  1183
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
  1184
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1185
fun simpfm :: "fm \<Rightarrow> fm"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1186
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
  1187
  "simpfm (And p q) = conj (simpfm p) (simpfm q)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1188
| "simpfm (Or p q) = disj (simpfm p) (simpfm q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1189
| "simpfm (Imp p q) = imp (simpfm p) (simpfm q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1190
| "simpfm (Iff p q) = iff (simpfm p) (simpfm q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1191
| "simpfm (NOT p) = not (simpfm p)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1192
| "simpfm (Lt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v < 0) then T else F
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1193
    | _ \<Rightarrow> Lt (reducecoeff a'))"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1194
| "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
  1195
| "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
  1196
| "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
  1197
| "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
  1198
| "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
  1199
| "simpfm (Dvd i a) = (if i=0 then simpfm (Eq a)
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1200
             else if (\<bar>i\<bar> = 1) \<and> check_int a then 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
  1201
             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))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1202
| "simpfm (NDvd i a) = (if i=0 then simpfm (NEq a)
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1203
             else if (\<bar>i\<bar> = 1) \<and> check_int a then 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
  1204
             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
  1205
| "simpfm 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
  1206
324622260d29 Added 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
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
  1208
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
  1209
  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
  1210
  {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
  1211
  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
  1212
    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
  1213
    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
  1214
    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
  1215
    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
  1216
    {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
  1217
    with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1218
    hence gp: "real_of_int ?g > 0" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1219
    have "Inum bs ?sa = real_of_int ?g* ?r" by (simp add: reducecoeff)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1220
    with sa have "Inum bs a < 0 = (real_of_int ?g * ?r < real_of_int ?g * 0)" 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
  1221
    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
  1222
      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
  1223
    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
  1224
  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
  1225
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
  1226
  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
  1227
  {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
  1228
  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
  1229
    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
  1230
    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
  1231
    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
  1232
    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
  1233
    {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
  1234
    with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1235
    hence gp: "real_of_int ?g > 0" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1236
    have "Inum bs ?sa = real_of_int ?g* ?r" by (simp add: reducecoeff)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1237
    with sa have "Inum bs a \<le> 0 = (real_of_int ?g * ?r \<le> real_of_int ?g * 0)" 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
  1238
    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
  1239
      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
  1240
    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
  1241
  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
  1242
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
  1243
  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
  1244
  {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
  1245
  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
  1246
    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
  1247
    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
  1248
    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
  1249
    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
  1250
    {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
  1251
    with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1252
    hence gp: "real_of_int ?g > 0" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1253
    have "Inum bs ?sa = real_of_int ?g* ?r" by (simp add: reducecoeff)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1254
    with sa have "Inum bs a > 0 = (real_of_int ?g * ?r > real_of_int ?g * 0)" 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
  1255
    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
  1256
      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
  1257
    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
  1258
  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
  1259
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
  1260
  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
  1261
  {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
  1262
  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
  1263
    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
  1264
    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
  1265
    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
  1266
    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
  1267
    {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
  1268
    with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1269
    hence gp: "real_of_int ?g > 0" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1270
    have "Inum bs ?sa = real_of_int ?g* ?r" by (simp add: reducecoeff)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1271
    with sa have "Inum bs a \<ge> 0 = (real_of_int ?g * ?r \<ge> real_of_int ?g * 0)" 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
  1272
    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
  1273
      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
  1274
    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
  1275
  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
  1276
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
  1277
  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
  1278
  {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
  1279
  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
  1280
    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
  1281
    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
  1282
    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
  1283
    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
  1284
    {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
  1285
    with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1286
    hence gp: "real_of_int ?g > 0" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1287
    have "Inum bs ?sa = real_of_int ?g* ?r" by (simp add: reducecoeff)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1288
    with sa have "Inum bs a = 0 = (real_of_int ?g * ?r = 0)" 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
  1289
    also have "\<dots> = (?r = 0)" using gp
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  1290
      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
  1291
    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
  1292
  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
  1293
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
  1294
  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
  1295
  {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
  1296
  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
  1297
    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
  1298
    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
  1299
    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
  1300
    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
  1301
    {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
  1302
    with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1303
    hence gp: "real_of_int ?g > 0" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1304
    have "Inum bs ?sa = real_of_int ?g* ?r" by (simp add: reducecoeff)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1305
    with sa have "Inum bs a \<noteq> 0 = (real_of_int ?g * ?r \<noteq> 0)" 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
  1306
    also have "\<dots> = (?r \<noteq> 0)" using gp
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  1307
      by simp
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  1308
    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
  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 (12 i a)  let ?sa = "simpnum a"   have sa: "Inum bs ?sa = Inum bs a" by simp
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1312
  have "i=0 \<or> (\<bar>i\<bar> = 1 \<and> check_int a) \<or> (i\<noteq>0 \<and> ((\<bar>i\<bar> \<noteq> 1) \<or> (\<not> check_int a)))" 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
  1313
  {assume "i=0" hence ?case using "12.hyps" by (simp add: rdvd_left_0_eq Let_def)}
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1314
  moreover
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1315
  {assume ai1: "\<bar>i\<bar> = 1" and ai: "check_int 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
  1316
    hence "i=1 \<or> i= - 1" by arith
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1317
    moreover {assume i1: "i = 1"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1318
      from rdvd_left1_int[OF check_int[OF ai, 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
  1319
      have ?case using i1 ai by simp }
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1320
    moreover {assume i1: "i = - 1"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
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}
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1325
  moreover
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1326
  {assume inz: "i\<noteq>0" and cond: "(\<bar>i\<bar> \<noteq> 1) \<or> (\<not> check_int 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
  1327
    {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1328
        by (cases "\<bar>i\<bar> = 1", auto simp add: int_rdvd_iff) }
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1329
    moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
23264
324622260d29 Added 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 (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
  1331
      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
  1332
      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
  1333
    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
  1334
  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
  1335
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
  1336
  case (13 i a)  let ?sa = "simpnum a"   have sa: "Inum bs ?sa = Inum bs a" by simp
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1337
  have "i=0 \<or> (\<bar>i\<bar> = 1 \<and> check_int a) \<or> (i\<noteq>0 \<and> ((\<bar>i\<bar> \<noteq> 1) \<or> (\<not> check_int a)))" 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
  1338
  {assume "i=0" hence ?case using "13.hyps" by (simp add: rdvd_left_0_eq Let_def)}
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1339
  moreover
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1340
  {assume ai1: "\<bar>i\<bar> = 1" and ai: "check_int 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
  1341
    hence "i=1 \<or> i= - 1" by arith
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1342
    moreover {assume i1: "i = 1"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1343
      from rdvd_left1_int[OF check_int[OF ai, 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
  1344
      have ?case using i1 ai by simp }
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1345
    moreover {assume i1: "i = - 1"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1346
      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
  1347
        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
  1348
      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
  1349
    ultimately have ?case by blast}
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1350
  moreover
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1351
  {assume inz: "i\<noteq>0" and cond: "(\<bar>i\<bar> \<noteq> 1) \<or> (\<not> check_int 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
  1352
    {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1353
        by (cases "\<bar>i\<bar> = 1", auto simp add: int_rdvd_iff) }
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1354
    moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1355
      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
  1356
        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
  1357
      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
  1358
      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
  1359
    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
  1360
  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
  1361
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
  1362
324622260d29 Added 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
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
  1364
  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
  1365
324622260d29 Added 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
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
  1367
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
  1368
  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
  1369
  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
  1370
  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
  1371
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
  1372
  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
  1373
  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
  1374
  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
  1375
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
  1376
  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
  1377
  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
  1378
  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
  1379
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
  1380
  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
  1381
  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
  1382
  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
  1383
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
  1384
  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
  1385
  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
  1386
  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
  1387
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
  1388
  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
  1389
  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
  1390
  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
  1391
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
  1392
  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
  1393
  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
  1394
  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
  1395
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
  1396
  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
  1397
  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
  1398
  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
  1399
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
  1400
324622260d29 Added 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
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
  1402
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
  1403
(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
  1404
324622260d29 Added 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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1406
  (* 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
  1407
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  1408
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
  1409
  "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
  1410
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
  1411
  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
  1412
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
  1413
  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
  1414
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
  1415
  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
  1416
definition CJNB :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm" where
29788
1b80ebe713a4 established session HOL-Reflection
haftmann
parents: 29667
diff changeset
  1417
  "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
  1418
                   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
  1419
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1420
lemma CJNB_qe:
23264
324622260d29 Added 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
  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
  1422
  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
  1423
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
  1424
  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
  1425
  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
  1426
  let ?cjs = "conjuncts p"
29788
1b80ebe713a4 established session HOL-Reflection
haftmann
parents: 29667
diff changeset
  1427
  let ?yes = "fst (List.partition bound0 ?cjs)"
1b80ebe713a4 established session HOL-Reflection
haftmann
parents: 29667
diff changeset
  1428
  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
  1429
  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
  1430
  let ?cyes = "list_conj ?yes"
29788
1b80ebe713a4 established session HOL-Reflection
haftmann
parents: 29667
diff changeset
  1431
  have part: "List.partition bound0 ?cjs = (?yes,?no)" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1432
  from partition_P[OF part] have "\<forall> q\<in> set ?yes. bound0 q" by blast
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1433
  hence yes_nb: "bound0 ?cyes" by (simp add: list_conj_nb)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1434
  hence yes_qf: "qfree (decr ?cyes )" by (simp add: decr_qf)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1435
  from conjuncts_qf[OF qfp] partition_set[OF 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
  1436
  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
  1437
  hence no_qf: "qfree ?cno"by (simp add: list_conj_qf)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1438
  with qe have cno_qf:"qfree (qe ?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
  1439
    and noE: "Ifm bs (qe ?cno) = Ifm bs (E ?cno)" by blast+
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1440
  from cno_qf yes_qf have qf: "qfree (CJNB qe p)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  1441
    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
  1442
  {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
  1443
    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
  1444
    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
  1445
      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
  1446
    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
  1447
  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
  1448
  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
  1449
    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
  1450
  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
  1451
    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
  1452
  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
  1453
    using qe[rule_format, OF no_qf] by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1454
  finally have "Ifm bs (E p) = Ifm bs (CJNB qe 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
  1455
    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
  1456
  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
  1457
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
  1458
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1459
fun qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1460
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
  1461
  "qelim (E p) = (\<lambda> qe. DJ (CJNB qe) (qelim p qe))"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1462
| "qelim (A p) = (\<lambda> qe. not (qe ((qelim (NOT p) qe))))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1463
| "qelim (NOT p) = (\<lambda> qe. not (qelim p qe))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1464
| "qelim (And p q) = (\<lambda> qe. conj (qelim p qe) (qelim q qe))"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1465
| "qelim (Or  p q) = (\<lambda> qe. disj (qelim p qe) (qelim q qe))"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1466
| "qelim (Imp p q) = (\<lambda> qe. disj (qelim (NOT p) qe) (qelim q qe))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1467
| "qelim (Iff p q) = (\<lambda> qe. iff (qelim p qe) (qelim q qe))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1468
| "qelim p = (\<lambda> y. simpfm 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
  1469
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1470
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
  1471
  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
  1472
  shows "\<And> bs. qfree (qelim p qe) \<and> (Ifm bs (qelim p qe) = Ifm bs p)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1473
  using qe_inv DJ_qe[OF CJNB_qe[OF qe_inv]]
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1474
  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
  1475
324622260d29 Added 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
61586
5197a2ecb658 isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
  1477
text \<open>The \<open>\<int>\<close> Part\<close>
5197a2ecb658 isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
  1478
text\<open>Linearity for fm where Bound 0 ranges over \<open>\<int>\<close>\<close>
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1479
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1480
fun zsplit0 :: "num \<Rightarrow> int \<times> num" (* splits the bounded from the unbounded part*)
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1481
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
  1482
  "zsplit0 (C c) = (0,C c)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1483
| "zsplit0 (Bound n) = (if n=0 then (1, C 0) else (0,Bound n))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1484
| "zsplit0 (CN n c a) = zsplit0 (Add (Mul c (Bound n)) a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1485
| "zsplit0 (CF c a b) = zsplit0 (Add (Mul c (Floor a)) b)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1486
| "zsplit0 (Neg a) = (let (i',a') =  zsplit0 a in (-i', Neg a'))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1487
| "zsplit0 (Add a b) = (let (ia,a') =  zsplit0 a ;
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1488
                            (ib,b') =  zsplit0 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
  1489
                            in (ia+ib, Add a' b'))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1490
| "zsplit0 (Sub a b) = (let (ia,a') =  zsplit0 a ;
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1491
                            (ib,b') =  zsplit0 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
  1492
                            in (ia-ib, Sub a' b'))"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1493
| "zsplit0 (Mul i a) = (let (i',a') =  zsplit0 a in (i*i', Mul i a'))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1494
| "zsplit0 (Floor a) = (let (i',a') =  zsplit0 a in (i',Floor 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
  1495
324622260d29 Added 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
lemma zsplit0_I:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1497
  shows "\<And> n a. zsplit0 t = (n,a) \<Longrightarrow> (Inum ((real_of_int (x::int)) #bs) (CN 0 n a) = Inum (real_of_int x #bs) t) \<and> numbound0 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
  1498
  (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
  1499
proof(induct t rule: zsplit0.induct)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1500
  case (1 c n a) thus ?case 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
  1501
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
  1502
  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
  1503
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
  1504
  case (3 n i a n a') thus ?case by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1505
next
23264
324622260d29 Added 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
  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
  1507
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
  1508
  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
  1509
  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
  1510
  let ?at = "snd (zsplit0 t)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1511
  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
  1512
    by (simp add: Let_def split_def)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1513
  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
  1514
  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
  1515
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
  1516
  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
  1517
  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
  1518
  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
  1519
  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
  1520
  let ?at = "snd (zsplit0 t)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1521
  have abjs: "zsplit0 s = (?ns,?as)" by simp
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1522
  moreover have abjt:  "zsplit0 t = (?nt,?at)" by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1523
  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
  1524
    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
  1525
  from abjs[symmetric] have bluddy: "\<exists> x y. (x,y) = zsplit0 s" by blast
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1526
  from 6 have "(\<exists> x y. (x,y) = zsplit0 s) \<longrightarrow> (\<forall>xa xb. zsplit0 t = (xa, xb) \<longrightarrow> Inum (real_of_int x # bs) (CN 0 xa xb) = Inum (real_of_int 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
  1527
  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
  1528
  from abjs 6  have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \<and> ?N ?as" by blast
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1529
  from th3[simplified] th2[simplified] th[simplified] show ?case
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 49069
diff changeset
  1530
    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
  1531
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
  1532
  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
  1533
  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
  1534
  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
  1535
  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
  1536
  let ?at = "snd (zsplit0 t)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1537
  have abjs: "zsplit0 s = (?ns,?as)" by simp
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1538
  moreover have abjt:  "zsplit0 t = (?nt,?at)" by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1539
  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
  1540
    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
  1541
  from abjs[symmetric] have bluddy: "\<exists> x y. (x,y) = zsplit0 s" by blast
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1542
  from 7 have "(\<exists> x y. (x,y) = zsplit0 s) \<longrightarrow> (\<forall>xa xb. zsplit0 t = (xa, xb) \<longrightarrow> Inum (real_of_int x # bs) (CN 0 xa xb) = Inum (real_of_int 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
  1543
  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
  1544
  from abjs 7 have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \<and> ?N ?as" by blast
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1545
  from th3[simplified] th2[simplified] th[simplified] show ?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
  1546
    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
  1547
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
  1548
  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
  1549
  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
  1550
  let ?at = "snd (zsplit0 t)"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1551
  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
  1552
    by (simp add: Let_def split_def)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1553
  from abj 8 have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1554
  hence "?I x (Mul i t) = (real_of_int i) * ?I x (CN 0 ?nt ?at)" by simp
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 49069
diff changeset
  1555
  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
  1556
  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
  1557
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
  1558
  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
  1559
  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
  1560
  let ?at = "snd (zsplit0 t)"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1561
  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
  1562
    by (simp add: Let_def split_def)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1563
  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
  1564
  hence na: "?N a" using th by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1565
  have th': "(real_of_int ?nt)*(real_of_int x) = real_of_int (?nt * x)" 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
  1566
  have "?I x (Floor t) = ?I x (Floor (CN 0 ?nt ?at))" using th2 by simp
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1567
  also have "\<dots> = real_of_int \<lfloor>real_of_int ?nt * real_of_int x + ?I x ?at\<rfloor>" by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1568
  also have "\<dots> = real_of_int \<lfloor>?I x ?at + real_of_int (?nt * x)\<rfloor>" by (simp add: ac_simps)
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1569
  also have "\<dots> = real_of_int (\<lfloor>?I x ?at\<rfloor> + (?nt * x))"
63600
d0fa16751d14 fixed floor proof
nipkow
parents: 62342
diff changeset
  1570
    by (simp add: of_int_mult[symmetric] del: of_int_mult)
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1571
  also have "\<dots> = real_of_int (?nt)*(real_of_int x) + real_of_int \<lfloor>?I x ?at\<rfloor>" 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
  1572
  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
  1573
  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
  1574
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
  1575
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1576
fun iszlfm :: "fm \<Rightarrow> real list \<Rightarrow> bool"   (* Linearity test for fm *)
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1577
where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1578
  "iszlfm (And p q) = (\<lambda> bs. iszlfm p bs \<and> iszlfm q bs)"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1579
| "iszlfm (Or p q) = (\<lambda> bs. iszlfm p bs \<and> iszlfm q bs)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1580
| "iszlfm (Eq  (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1581
| "iszlfm (NEq (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1582
| "iszlfm (Lt  (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1583
| "iszlfm (Le  (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1584
| "iszlfm (Gt  (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1585
| "iszlfm (Ge  (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1586
| "iszlfm (Dvd i (CN 0 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
  1587
                 (\<lambda> bs. c>0 \<and> i>0 \<and> numbound0 e \<and> isint e bs)"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1588
| "iszlfm (NDvd i (CN 0 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
  1589
                 (\<lambda> bs. c>0 \<and> i>0 \<and> numbound0 e \<and> isint e bs)"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1590
| "iszlfm p = (\<lambda> bs. isatom p \<and> (bound0 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
  1591
324622260d29 Added 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
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
  1593
  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
  1594
324622260d29 Added 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
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
  1596
  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
  1597
  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
  1598
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
  1599
  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
  1600
  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
  1601
    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
  1602
  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
  1603
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
  1604
324622260d29 Added 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
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
  1606
  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
  1607
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
  1608
  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
  1609
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1610
fun zlfm :: "fm \<Rightarrow> fm"       (* Linearity transformation for fm *)
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1611
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
  1612
  "zlfm (And p q) = conj (zlfm p) (zlfm q)"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1613
| "zlfm (Or p q) = disj (zlfm p) (zlfm q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1614
| "zlfm (Imp p q) = disj (zlfm (NOT p)) (zlfm q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1615
| "zlfm (Iff p q) = disj (conj (zlfm p) (zlfm q)) (conj (zlfm (NOT p)) (zlfm (NOT q)))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1616
| "zlfm (Lt a) = (let (c,r) = zsplit0 a in
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1617
     if c=0 then Lt r else
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1618
     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)))
23264
324622260d29 Added 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 Or (Gt (CN 0 (-c) (Floor(Neg r)))) (And (Eq(CN 0 (-c) (Floor(Neg r)))) (Lt (Add (Floor (Neg r)) r))))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1620
| "zlfm (Le a) = (let (c,r) = zsplit0 a in
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1621
     if c=0 then Le r else
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1622
     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)))
23264
324622260d29 Added 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
     else Or (Ge (CN 0 (-c) (Floor(Neg r)))) (And (Eq(CN 0 (-c) (Floor(Neg r)))) (Lt (Add (Floor (Neg r)) r))))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1624
| "zlfm (Gt a) = (let (c,r) = zsplit0 a in
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1625
     if c=0 then Gt r else
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1626
     if c>0 then Or (Gt (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) 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
  1627
     else Or (Lt (CN 0 (-c) (Neg (Floor r)))) (And (Eq(CN 0 (-c) (Neg (Floor r)))) (Lt (Sub (Floor r) r))))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1628
| "zlfm (Ge a) = (let (c,r) = zsplit0 a in
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1629
     if c=0 then Ge r else
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1630
     if c>0 then Or (Ge (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) 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
  1631
     else Or (Le (CN 0 (-c) (Neg (Floor r)))) (And (Eq(CN 0 (-c) (Neg (Floor r)))) (Lt (Sub (Floor r) r))))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1632
| "zlfm (Eq a) = (let (c,r) = zsplit0 a in
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1633
              if c=0 then Eq r 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
  1634
      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
  1635
      else (And (Eq (CN 0 (-c) (Floor (Neg r)))) (Eq (Add (Floor (Neg r)) r))))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1636
| "zlfm (NEq a) = (let (c,r) = zsplit0 a in
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1637
              if c=0 then NEq r 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
  1638
      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
  1639
      else (Or (NEq (CN 0 (-c) (Floor (Neg r)))) (NEq (Add (Floor (Neg r)) r))))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1640
| "zlfm (Dvd i a) = (if i=0 then zlfm (Eq a)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1641
  else (let (c,r) = zsplit0 a in
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1642
              if c=0 then Dvd \<bar>i\<bar> r else
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1643
      if c>0 then And (Eq (Sub (Floor r) r)) (Dvd \<bar>i\<bar> (CN 0 c (Floor r)))
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1644
      else And (Eq (Sub (Floor r) r)) (Dvd \<bar>i\<bar> (CN 0 (-c) (Neg (Floor r))))))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1645
| "zlfm (NDvd i a) = (if i=0 then zlfm (NEq a)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1646
  else (let (c,r) = zsplit0 a in
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1647
              if c=0 then NDvd \<bar>i\<bar> r else
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1648
      if c>0 then Or (NEq (Sub (Floor r) r)) (NDvd \<bar>i\<bar> (CN 0 c (Floor r)))
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1649
      else Or (NEq (Sub (Floor r) r)) (NDvd \<bar>i\<bar> (CN 0 (-c) (Neg (Floor r))))))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1650
| "zlfm (NOT (And p q)) = disj (zlfm (NOT p)) (zlfm (NOT q))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1651
| "zlfm (NOT (Or p q)) = conj (zlfm (NOT p)) (zlfm (NOT q))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1652
| "zlfm (NOT (Imp p q)) = conj (zlfm p) (zlfm (NOT q))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1653
| "zlfm (NOT (Iff p q)) = disj (conj(zlfm p) (zlfm(NOT q))) (conj (zlfm(NOT p)) (zlfm q))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1654
| "zlfm (NOT (NOT p)) = zlfm p"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1655
| "zlfm (NOT T) = F"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1656
| "zlfm (NOT F) = T"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1657
| "zlfm (NOT (Lt a)) = zlfm (Ge a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1658
| "zlfm (NOT (Le a)) = zlfm (Gt a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1659
| "zlfm (NOT (Gt a)) = zlfm (Le a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1660
| "zlfm (NOT (Ge a)) = zlfm (Lt a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1661
| "zlfm (NOT (Eq a)) = zlfm (NEq a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1662
| "zlfm (NOT (NEq a)) = zlfm (Eq a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1663
| "zlfm (NOT (Dvd i a)) = zlfm (NDvd i a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1664
| "zlfm (NOT (NDvd i a)) = zlfm (Dvd i a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1665
| "zlfm 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
  1666
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1667
lemma split_int_less_real:
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1668
  "(real_of_int (a::int) < b) = (a < \<lfloor>b\<rfloor> \<or> (a = \<lfloor>b\<rfloor> \<and> real_of_int \<lfloor>b\<rfloor> < 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
proof( auto)
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1670
  assume alb: "real_of_int a < b" and agb: "\<not> a < \<lfloor>b\<rfloor>"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1671
  from agb have "\<lfloor>b\<rfloor> \<le> a" by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1672
  hence th: "b < real_of_int a + 1" by (simp only: floor_le_iff)
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1673
  from floor_eq[OF alb th] show "a = \<lfloor>b\<rfloor>" 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
  1674
next
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1675
  assume alb: "a < \<lfloor>b\<rfloor>"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1676
  hence "real_of_int a < real_of_int \<lfloor>b\<rfloor>" by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1677
  moreover have "real_of_int \<lfloor>b\<rfloor> \<le> b" by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1678
  ultimately show  "real_of_int a < b" 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
  1679
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
  1680
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1681
lemma split_int_less_real':
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1682
  "(real_of_int (a::int) + b < 0) = (real_of_int a - real_of_int \<lfloor>- b\<rfloor> < 0 \<or> (real_of_int a - real_of_int \<lfloor>- b\<rfloor> = 0 \<and> real_of_int \<lfloor>- b\<rfloor> + b < 0))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1683
proof-
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1684
  have "(real_of_int a + b <0) = (real_of_int a < -b)" by arith
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1685
  with split_int_less_real[where a="a" and b="-b"] show ?thesis 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
  1686
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
  1687
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1688
lemma split_int_gt_real':
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1689
  "(real_of_int (a::int) + b > 0) = (real_of_int a + real_of_int \<lfloor>b\<rfloor> > 0 \<or> (real_of_int a + real_of_int \<lfloor>b\<rfloor> = 0 \<and> real_of_int \<lfloor>b\<rfloor> - b < 0))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1690
proof-
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1691
  have th: "(real_of_int a + b >0) = (real_of_int (-a) + (-b)< 0)" by arith
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  1692
  show ?thesis 
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  1693
    by (simp only:th split_int_less_real'[where a="-a" and b="-b"]) (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
  1694
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
  1695
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1696
lemma split_int_le_real:
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1697
  "(real_of_int (a::int) \<le> b) = (a \<le> \<lfloor>b\<rfloor> \<or> (a = \<lfloor>b\<rfloor> \<and> real_of_int \<lfloor>b\<rfloor> < 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
  1698
proof( auto)
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1699
  assume alb: "real_of_int a \<le> b" and agb: "\<not> a \<le> \<lfloor>b\<rfloor>"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1700
  from alb have "\<lfloor>real_of_int a\<rfloor> \<le> \<lfloor>b\<rfloor>" by (simp only: floor_mono)
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1701
  hence "a \<le> \<lfloor>b\<rfloor>" by simp with agb show "False" 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
  1702
next
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1703
  assume alb: "a \<le> \<lfloor>b\<rfloor>"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1704
  hence "real_of_int a \<le> real_of_int \<lfloor>b\<rfloor>" by (simp only: floor_mono)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1705
  also have "\<dots>\<le> b" by simp  finally show  "real_of_int a \<le> 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
  1706
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
  1707
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1708
lemma split_int_le_real':
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1709
  "(real_of_int (a::int) + b \<le> 0) = (real_of_int a - real_of_int \<lfloor>- b\<rfloor> \<le> 0 \<or> (real_of_int a - real_of_int \<lfloor>- b\<rfloor> = 0 \<and> real_of_int \<lfloor>- b\<rfloor> + b < 0))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1710
proof-
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1711
  have "(real_of_int a + b \<le>0) = (real_of_int a \<le> -b)" by arith
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1712
  with split_int_le_real[where a="a" and b="-b"] show ?thesis 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
  1713
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
  1714
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1715
lemma split_int_ge_real':
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1716
  "(real_of_int (a::int) + b \<ge> 0) = (real_of_int a + real_of_int \<lfloor>b\<rfloor> \<ge> 0 \<or> (real_of_int a + real_of_int \<lfloor>b\<rfloor> = 0 \<and> real_of_int \<lfloor>b\<rfloor> - b < 0))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1717
proof-
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1718
  have th: "(real_of_int a + b \<ge>0) = (real_of_int (-a) + (-b) \<le> 0)" 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
  1719
  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
  1720
    (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
  1721
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
  1722
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1723
lemma split_int_eq_real: "(real_of_int (a::int) = b) = ( a = \<lfloor>b\<rfloor> \<and> b = real_of_int \<lfloor>b\<rfloor>)" (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
  1724
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
  1725
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1726
lemma split_int_eq_real': "(real_of_int (a::int) + b = 0) = ( a - \<lfloor>- b\<rfloor> = 0 \<and> real_of_int \<lfloor>- b\<rfloor> + b = 0)" (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
  1727
proof-
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1728
  have "?l = (real_of_int a = -b)" 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
  1729
  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
  1730
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
  1731
324622260d29 Added 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
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
  1733
  assumes qfp: "qfree p"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1734
  shows "(Ifm (real_of_int i #bs) (zlfm p) = Ifm (real_of_int i# bs) p) \<and> iszlfm (zlfm p) (real_of_int (i::int) #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
  1735
  (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
  1736
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
  1737
proof(induct p rule: zlfm.induct)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1738
  case (5 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
  1739
  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
  1740
  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
  1741
  have spl: "zsplit0 a = (?c,?r)" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1742
  from zsplit0_I[OF spl, where x="i" and bs="bs"]
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1743
  have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1744
  let ?N = "\<lambda> t. Inum (real_of_int i#bs) 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
  1745
  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
  1746
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1747
  {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"]
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1748
      by (cases "?r", simp_all add: Let_def split_def,rename_tac nat a b,case_tac "nat", 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
  1749
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1750
  {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Lt 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
  1751
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1752
    have "?I (Lt a) = (real_of_int (?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
  1753
    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
  1754
    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
  1755
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1756
  {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Lt 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
  1757
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1758
    have "?I (Lt a) = (real_of_int (?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
  1759
    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
  1760
    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
  1761
  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
  1762
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
  1763
  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
  1764
  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
  1765
  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
  1766
  have spl: "zsplit0 a = (?c,?r)" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1767
  from zsplit0_I[OF spl, where x="i" and bs="bs"]
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1768
  have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1769
  let ?N = "\<lambda> t. Inum (real_of_int i#bs) 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
  1770
  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
  1771
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1772
  {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"]
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1773
      by (cases "?r", simp_all add: Let_def split_def, rename_tac nat a b, case_tac "nat",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
  1774
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1775
  {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (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
  1776
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1777
    have "?I (Le a) = (real_of_int (?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
  1778
    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
  1779
    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
  1780
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1781
  {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (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
  1782
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1783
    have "?I (Le a) = (real_of_int (?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
  1784
    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
  1785
    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
  1786
  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
  1787
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1788
  case (7 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
  1789
  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
  1790
  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
  1791
  have spl: "zsplit0 a = (?c,?r)" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1792
  from zsplit0_I[OF spl, where x="i" and bs="bs"]
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1793
  have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1794
  let ?N = "\<lambda> t. Inum (real_of_int i#bs) 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
  1795
  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
  1796
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1797
  {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"]
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1798
      by (cases "?r", simp_all add: Let_def split_def, rename_tac nat a b, case_tac "nat", 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
  1799
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1800
  {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Gt 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
  1801
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1802
    have "?I (Gt a) = (real_of_int (?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
  1803
    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
  1804
    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
  1805
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1806
  {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Gt 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
  1807
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1808
    have "?I (Gt a) = (real_of_int (?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
  1809
    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
  1810
    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
  1811
  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
  1812
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
  1813
  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
  1814
   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
  1815
  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
  1816
  have spl: "zsplit0 a = (?c,?r)" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1817
  from zsplit0_I[OF spl, where x="i" and bs="bs"]
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1818
  have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1819
  let ?N = "\<lambda> t. Inum (real_of_int i#bs) 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
  1820
  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
  1821
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1822
  {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"]
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1823
      by (cases "?r", simp_all add: Let_def split_def, rename_tac nat a b, case_tac "nat", 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
  1824
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1825
  {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (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
  1826
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1827
    have "?I (Ge a) = (real_of_int (?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
  1828
    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
  1829
    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
  1830
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1831
  {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (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
  1832
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1833
    have "?I (Ge a) = (real_of_int (?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
  1834
    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
  1835
    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
  1836
  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
  1837
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
  1838
  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
  1839
  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
  1840
  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
  1841
  have spl: "zsplit0 a = (?c,?r)" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1842
  from zsplit0_I[OF spl, where x="i" and bs="bs"]
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1843
  have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1844
  let ?N = "\<lambda> t. Inum (real_of_int i#bs) 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
  1845
  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
  1846
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1847
  {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"]
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1848
      by (cases "?r", simp_all add: Let_def split_def, rename_tac nat a b, case_tac "nat", 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
  1849
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1850
  {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (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
  1851
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1852
    have "?I (Eq a) = (real_of_int (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1853
    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 of_int_mult[symmetric] del: 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
  1854
    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
  1855
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1856
  {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (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
  1857
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1858
    have "?I (Eq a) = (real_of_int (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1859
    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 of_int_mult[symmetric] del: of_int_mult,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
  1860
    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
  1861
  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
  1862
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
  1863
  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
  1864
  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
  1865
  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
  1866
  have spl: "zsplit0 a = (?c,?r)" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1867
  from zsplit0_I[OF spl, where x="i" and bs="bs"]
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1868
  have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1869
  let ?N = "\<lambda> t. Inum (real_of_int i#bs) 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
  1870
  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
  1871
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1872
  {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"]
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1873
      by (cases "?r", simp_all add: Let_def split_def, rename_tac nat a b, case_tac "nat", 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
  1874
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1875
  {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (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
  1876
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1877
    have "?I (NEq a) = (real_of_int (?c * i) + (?N ?r) \<noteq> 0)" using Ia by (simp add: Let_def split_def)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1878
    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 of_int_mult[symmetric] del: 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
  1879
    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
  1880
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1881
  {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (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
  1882
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1883
    have "?I (NEq a) = (real_of_int (?c * i) + (?N ?r) \<noteq> 0)" using Ia by (simp add: Let_def split_def)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1884
    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 of_int_mult[symmetric] del: of_int_mult,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
  1885
    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
  1886
  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
  1887
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
  1888
  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
  1889
  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
  1890
  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
  1891
  have spl: "zsplit0 a = (?c,?r)" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1892
  from zsplit0_I[OF spl, where x="i" and bs="bs"]
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1893
  have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1894
  let ?N = "\<lambda> t. Inum (real_of_int i#bs) 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
  1895
  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
  1896
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1897
  { assume j: "j=0" hence z: "zlfm (Dvd j a) = (zlfm (Eq a))" by (simp add: Let_def)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1898
    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
  1899
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1900
  {assume "?c=0" and "j\<noteq>0" hence ?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
  1901
      using zsplit0_I[OF spl, where x="i" and bs="bs"] rdvd_abs1[where d="j"]
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1902
      by (cases "?r", simp_all add: Let_def split_def, rename_tac nat a b, case_tac "nat", 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
  1903
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1904
  {assume cp: "?c > 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (Dvd j 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
  1905
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1906
    have "?I (Dvd j a) = (real_of_int j rdvd (real_of_int (?c * i) + (?N ?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
  1907
      using Ia by (simp add: Let_def split_def)
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1908
    also have "\<dots> = (real_of_int \<bar>j\<bar> rdvd real_of_int (?c*i) + (?N ?r))"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1909
      by (simp only: rdvd_abs1[where d="j" and t="real_of_int (?c*i) + ?N ?r", symmetric]) simp
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1910
    also have "\<dots> = (\<bar>j\<bar> dvd \<lfloor>(?N ?r) + real_of_int (?c*i)\<rfloor> \<and>
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1911
       (real_of_int \<lfloor>(?N ?r) + real_of_int (?c*i)\<rfloor> = (real_of_int (?c*i) + (?N ?r))))"
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1912
      by(simp only: int_rdvd_real[where i="\<bar>j\<bar>" and x="real_of_int (?c*i) + (?N ?r)"]) (simp only: ac_simps)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1913
    also have "\<dots> = (?I (?l (Dvd j a)))" using cp cnz jnz
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1914
      by (simp add: Let_def split_def int_rdvd_iff[symmetric]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1915
        del: 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
  1916
    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
  1917
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1918
  {assume cn: "?c < 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (Dvd j 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
  1919
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1920
    have "?I (Dvd j a) = (real_of_int j rdvd (real_of_int (?c * i) + (?N ?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
  1921
      using Ia by (simp add: Let_def split_def)
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1922
    also have "\<dots> = (real_of_int \<bar>j\<bar> rdvd real_of_int (?c*i) + (?N ?r))"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1923
      by (simp only: rdvd_abs1[where d="j" and t="real_of_int (?c*i) + ?N ?r", symmetric]) simp
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1924
    also have "\<dots> = (\<bar>j\<bar> dvd \<lfloor>(?N ?r) + real_of_int (?c*i)\<rfloor> \<and>
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1925
       (real_of_int \<lfloor>(?N ?r) + real_of_int (?c*i)\<rfloor> = (real_of_int (?c*i) + (?N ?r))))"
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1926
      by(simp only: int_rdvd_real[where i="\<bar>j\<bar>" and x="real_of_int (?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
  1927
    also have "\<dots> = (?I (?l (Dvd j a)))" using cn cnz jnz
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1928
      using rdvd_minus [where d="\<bar>j\<bar>" and t="real_of_int (?c*i + \<lfloor>?N ?r\<rfloor>)", simplified, symmetric]
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1929
      by (simp add: Let_def split_def int_rdvd_iff[symmetric]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1930
        del: 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
  1931
    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
  1932
  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
  1933
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
  1934
  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
  1935
  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
  1936
  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
  1937
  have spl: "zsplit0 a = (?c,?r)" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1938
  from zsplit0_I[OF spl, where x="i" and bs="bs"]
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1939
  have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1940
  let ?N = "\<lambda> t. Inum (real_of_int i#bs) 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
  1941
  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
  1942
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1943
  {assume j: "j=0" hence z: "zlfm (NDvd j a) = (zlfm (NEq a))" by (simp add: Let_def)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1944
    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
  1945
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1946
  {assume "?c=0" and "j\<noteq>0" hence ?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
  1947
      using zsplit0_I[OF spl, where x="i" and bs="bs"] rdvd_abs1[where d="j"]
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1948
      by (cases "?r", simp_all add: Let_def split_def, rename_tac nat a b, case_tac "nat", 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
  1949
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1950
  {assume cp: "?c > 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (NDvd j 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
  1951
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1952
    have "?I (NDvd j a) = (\<not> (real_of_int j rdvd (real_of_int (?c * i) + (?N ?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
  1953
      using Ia by (simp add: Let_def split_def)
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1954
    also have "\<dots> = (\<not> (real_of_int \<bar>j\<bar> rdvd real_of_int (?c*i) + (?N ?r)))"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1955
      by (simp only: rdvd_abs1[where d="j" and t="real_of_int (?c*i) + ?N ?r", symmetric]) simp
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1956
    also have "\<dots> = (\<not> (\<bar>j\<bar> dvd \<lfloor>(?N ?r) + real_of_int (?c*i)\<rfloor> \<and>
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1957
       (real_of_int \<lfloor>(?N ?r) + real_of_int (?c*i)\<rfloor> = (real_of_int (?c*i) + (?N ?r)))))"
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1958
      by(simp only: int_rdvd_real[where i="\<bar>j\<bar>" and x="real_of_int (?c*i) + (?N ?r)"]) (simp only: ac_simps)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1959
    also have "\<dots> = (?I (?l (NDvd j a)))" using cp cnz jnz
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1960
      by (simp add: Let_def split_def int_rdvd_iff[symmetric]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1961
        del: 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
  1962
    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
  1963
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1964
  {assume cn: "?c < 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (NDvd j 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
  1965
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1966
    have "?I (NDvd j a) = (\<not> (real_of_int j rdvd (real_of_int (?c * i) + (?N ?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
  1967
      using Ia by (simp add: Let_def split_def)
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1968
    also have "\<dots> = (\<not> (real_of_int \<bar>j\<bar> rdvd real_of_int (?c*i) + (?N ?r)))"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1969
      by (simp only: rdvd_abs1[where d="j" and t="real_of_int (?c*i) + ?N ?r", symmetric]) simp
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1970
    also have "\<dots> = (\<not> (\<bar>j\<bar> dvd \<lfloor>(?N ?r) + real_of_int (?c*i)\<rfloor> \<and>
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  1971
       (real_of_int \<lfloor>(?N ?r) + real_of_int (?c*i)\<rfloor> = (real_of_int (?c*i) + (?N ?r)))))"
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1972
      by(simp only: int_rdvd_real[where i="\<bar>j\<bar>" and x="real_of_int (?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
  1973
    also have "\<dots> = (?I (?l (NDvd j a)))" using cn cnz jnz
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1974
      using rdvd_minus [where d="\<bar>j\<bar>" and t="real_of_int (?c*i + \<lfloor>?N ?r\<rfloor>)", simplified, symmetric]
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1975
      by (simp add: Let_def split_def int_rdvd_iff[symmetric]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1976
        del: 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
  1977
    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
  1978
  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
  1979
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
  1980
61586
5197a2ecb658 isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
  1981
text\<open>plusinf : Virtual substitution of \<open>+\<infinity>\<close>
5197a2ecb658 isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
  1982
       minusinf: Virtual substitution of \<open>-\<infinity>\<close>
5197a2ecb658 isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
  1983
       \<open>\<delta>\<close> Compute lcm \<open>d| Dvd d  c*x+t \<in> p\<close>
5197a2ecb658 isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
  1984
       \<open>d_\<delta>\<close> checks if a given l divides all the ds above\<close>
23316
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  1985
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1986
fun minusinf:: "fm \<Rightarrow> fm"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  1987
where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1988
  "minusinf (And p q) = conj (minusinf p) (minusinf q)"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  1989
| "minusinf (Or p q) = disj (minusinf p) (minusinf q)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1990
| "minusinf (Eq  (CN 0 c e)) = F"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1991
| "minusinf (NEq (CN 0 c e)) = T"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1992
| "minusinf (Lt  (CN 0 c e)) = T"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1993
| "minusinf (Le  (CN 0 c e)) = T"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1994
| "minusinf (Gt  (CN 0 c e)) = F"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1995
| "minusinf (Ge  (CN 0 c e)) = F"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1996
| "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
  1997
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1998
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
  1999
  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
  2000
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2001
fun plusinf:: "fm \<Rightarrow> fm"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2002
where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2003
  "plusinf (And p q) = conj (plusinf p) (plusinf q)"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2004
| "plusinf (Or p q) = disj (plusinf p) (plusinf q)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  2005
| "plusinf (Eq  (CN 0 c e)) = F"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  2006
| "plusinf (NEq (CN 0 c e)) = T"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  2007
| "plusinf (Lt  (CN 0 c e)) = F"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  2008
| "plusinf (Le  (CN 0 c e)) = F"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  2009
| "plusinf (Gt  (CN 0 c e)) = T"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  2010
| "plusinf (Ge  (CN 0 c e)) = T"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  2011
| "plusinf p = p"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  2012
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2013
fun \<delta> :: "fm \<Rightarrow> int"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2014
where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2015
  "\<delta> (And p q) = lcm (\<delta> p) (\<delta> q)"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2016
| "\<delta> (Or p q) = lcm (\<delta> p) (\<delta> q)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  2017
| "\<delta> (Dvd i (CN 0 c e)) = i"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  2018
| "\<delta> (NDvd i (CN 0 c e)) = i"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  2019
| "\<delta> p = 1"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  2020
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2021
fun d_\<delta> :: "fm \<Rightarrow> int \<Rightarrow> bool"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2022
where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2023
  "d_\<delta> (And p q) = (\<lambda> d. d_\<delta> p d \<and> d_\<delta> q d)"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2024
| "d_\<delta> (Or p q) = (\<lambda> d. d_\<delta> p d \<and> d_\<delta> q d)"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2025
| "d_\<delta> (Dvd i (CN 0 c e)) = (\<lambda> d. i dvd d)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2026
| "d_\<delta> (NDvd i (CN 0 c e)) = (\<lambda> d. i dvd d)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2027
| "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
  2028
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2029
lemma delta_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
  2030
  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
  2031
  and d: "d dvd d'"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2032
  and ad: "d_\<delta> p d"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2033
  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
  2034
  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
  2035
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
  2036
  case (9 i c e)  thus ?case using d
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
  2037
    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
  2038
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
  2039
  case (10 i c e) thus ?case using d
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
  2040
    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
  2041
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
  2042
324622260d29 Added 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
lemma \<delta> : assumes lin:"iszlfm p bs"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2044
  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
  2045
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
  2046
proof (induct p rule: iszlfm.induct)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2047
  case (1 p 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
  2048
  let ?d = "\<delta> (And p q)"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2049
  from 1 lcm_pos_int have dp: "?d >0" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2050
  have d1: "\<delta> p dvd \<delta> (And p q)" using 1 by simp
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2051
  hence th: "d_\<delta> p ?d"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2052
    using delta_mono 1 by (simp only: iszlfm.simps) blast
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2053
  have "\<delta> q dvd \<delta> (And p q)" using 1 by simp
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2054
  hence th': "d_\<delta> q ?d" using delta_mono 1 by (simp only: iszlfm.simps) blast
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2055
  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
  2056
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2057
  case (2 p 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
  2058
  let ?d = "\<delta> (And p q)"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2059
  from 2 lcm_pos_int have dp: "?d >0" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2060
  have "\<delta> p dvd \<delta> (And p q)" using 2 by simp
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2061
  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
  2062
  have "\<delta> q dvd \<delta> (And p q)" using 2 by simp
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2063
  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
  2064
  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
  2065
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
  2066
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2067
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2068
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
  2069
  assumes linp: "iszlfm p (a # bs)"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2070
  shows "\<exists> (z::int). \<forall> x < z. Ifm ((real_of_int x)#bs) (minusinf p) = Ifm ((real_of_int 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
  2071
  (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
  2072
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
  2073
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
  2074
  case (1 f g)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2075
  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
  2076
  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
  2077
  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
  2078
  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
  2079
  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
  2080
  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
  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
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2083
  case (2 f g)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2084
  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
  2085
  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
  2086
  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
  2087
  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
  2088
  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
  2089
  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
  2090
  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
  2091
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2092
  case (3 c e)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2093
  then have "c > 0" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2094
  hence rcpos: "real_of_int c > 0" by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2095
  from 3 have nbe: "numbound0 e" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  2096
  fix y
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2097
  have "\<forall> x < \<lfloor>- (Inum (y#bs) e) / (real_of_int c)\<rfloor>. ?I x (?M (Eq (CN 0 c e))) = ?I x (Eq (CN 0 c e))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2098
  proof (simp add: less_floor_iff , rule allI, rule impI)
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2099
    fix x :: int
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2100
    assume A: "real_of_int x + 1 \<le> - (Inum (y # bs) e / real_of_int c)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2101
    hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2102
    with rcpos  have "(real_of_int c)*(real_of_int  x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))"
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36531
diff changeset
  2103
      by (simp only: mult_strict_left_mono [OF th1 rcpos])
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2104
    hence "real_of_int c * real_of_int x + Inum (y # bs) e \<noteq> 0"using rcpos  by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2105
    thus "real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \<noteq> 0"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2106
      using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"]  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
  2107
  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
  2108
  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
  2109
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2110
  case (4 c e)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2111
  then have "c > 0" by simp hence rcpos: "real_of_int c > 0" by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2112
  from 4 have nbe: "numbound0 e" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  2113
  fix y
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2114
  have "\<forall> x < \<lfloor>- (Inum (y#bs) e) / (real_of_int c)\<rfloor>. ?I x (?M (NEq (CN 0 c e))) = ?I x (NEq (CN 0 c e))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2115
  proof (simp add: less_floor_iff , rule allI, rule impI)
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2116
    fix x :: int
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2117
    assume A: "real_of_int x + 1 \<le> - (Inum (y # bs) e / real_of_int c)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2118
    hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2119
    with rcpos  have "(real_of_int c)*(real_of_int x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))"
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36531
diff changeset
  2120
      by (simp only: mult_strict_left_mono [OF th1 rcpos])
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2121
    hence "real_of_int c * real_of_int x + Inum (y # bs) e \<noteq> 0"using rcpos  by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2122
    thus "real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \<noteq> 0"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2123
      using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"]  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
  2124
  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
  2125
  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
  2126
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2127
  case (5 c e)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2128
  then have "c > 0" by simp hence rcpos: "real_of_int c > 0" by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2129
  from 5 have nbe: "numbound0 e" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  2130
  fix y
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2131
  have "\<forall> x < \<lfloor>- (Inum (y#bs) e) / (real_of_int c)\<rfloor>. ?I x (?M (Lt (CN 0 c e))) = ?I x (Lt (CN 0 c e))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2132
  proof (simp add: less_floor_iff , rule allI, rule impI)
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2133
    fix x :: int
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2134
    assume A: "real_of_int x + 1 \<le> - (Inum (y # bs) e / real_of_int c)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2135
    hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2136
    with rcpos  have "(real_of_int c)*(real_of_int x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))"
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36531
diff changeset
  2137
      by (simp only: mult_strict_left_mono [OF th1 rcpos])
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2138
    thus "real_of_int c * real_of_int x + Inum (real_of_int x # bs) e < 0"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2139
      using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"] rcpos 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
  2140
  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
  2141
  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
  2142
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2143
  case (6 c e)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2144
  then have "c > 0" by simp hence rcpos: "real_of_int c > 0" by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2145
  from 6 have nbe: "numbound0 e" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  2146
  fix y
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2147
  have "\<forall> x < \<lfloor>- (Inum (y#bs) e) / (real_of_int c)\<rfloor>. ?I x (?M (Le (CN 0 c e))) = ?I x (Le (CN 0 c e))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2148
  proof (simp add: less_floor_iff , rule allI, rule impI)
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2149
    fix x :: int
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2150
    assume A: "real_of_int x + 1 \<le> - (Inum (y # bs) e / real_of_int c)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2151
    hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2152
    with rcpos  have "(real_of_int c)*(real_of_int x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))"
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36531
diff changeset
  2153
      by (simp only: mult_strict_left_mono [OF th1 rcpos])
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2154
    thus "real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \<le> 0"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2155
      using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"] rcpos 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
  2156
  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
  2157
  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
  2158
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2159
  case (7 c e)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2160
  then have "c > 0" by simp hence rcpos: "real_of_int c > 0" by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2161
  from 7 have nbe: "numbound0 e" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  2162
  fix y
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2163
  have "\<forall> x < \<lfloor>- (Inum (y#bs) e) / (real_of_int c)\<rfloor>. ?I x (?M (Gt (CN 0 c e))) = ?I x (Gt (CN 0 c e))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2164
  proof (simp add: less_floor_iff , rule allI, rule impI)
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2165
    fix x :: int
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2166
    assume A: "real_of_int x + 1 \<le> - (Inum (y # bs) e / real_of_int c)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2167
    hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2168
    with rcpos  have "(real_of_int c)*(real_of_int x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))"
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36531
diff changeset
  2169
      by (simp only: mult_strict_left_mono [OF th1 rcpos])
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2170
    thus "\<not> (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e>0)"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2171
      using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"] rcpos 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
  2172
  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
  2173
  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
  2174
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2175
  case (8 c e)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2176
  then have "c > 0" by simp hence rcpos: "real_of_int c > 0" by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2177
  from 8 have nbe: "numbound0 e" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  2178
  fix y
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2179
  have "\<forall> x < \<lfloor>- (Inum (y#bs) e) / (real_of_int c)\<rfloor>. ?I x (?M (Ge (CN 0 c e))) = ?I x (Ge (CN 0 c e))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2180
  proof (simp add: less_floor_iff , rule allI, rule impI)
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2181
    fix x :: int
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2182
    assume A: "real_of_int x + 1 \<le> - (Inum (y # bs) e / real_of_int c)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2183
    hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2184
    with rcpos  have "(real_of_int c)*(real_of_int x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))"
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36531
diff changeset
  2185
      by (simp only: mult_strict_left_mono [OF th1 rcpos])
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2186
    thus "\<not> real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \<ge> 0"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2187
      using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"] rcpos 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
  2188
  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
  2189
  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
  2190
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
  2191
324622260d29 Added 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
lemma minusinf_repeats:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2193
  assumes d: "d_\<delta> p d" and linp: "iszlfm p (a # bs)"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2194
  shows "Ifm ((real_of_int(x - k*d))#bs) (minusinf p) = Ifm (real_of_int x #bs) (minusinf 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
  2195
using linp d
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2196
proof(induct p rule: iszlfm.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
  2197
  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
  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
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2200
    show ?case
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2201
    proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="real_of_int x - real_of_int k * real_of_int d" and b'="real_of_int x"] right_diff_distrib, rule iffI)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2202
      assume
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2203
        "real_of_int i rdvd real_of_int c * real_of_int x - real_of_int c * (real_of_int k * real_of_int d) + Inum (real_of_int 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")
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2205
      hence "\<exists> (l::int). ?rt = ?ri * (real_of_int l)" by (simp add: rdvd_def)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2206
      hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real_of_int l)+?rc*(?rk * (real_of_int i) * (real_of_int 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)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2208
      hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real_of_int (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)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2210
      hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri* (real_of_int l)" by blast
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2211
      thus "real_of_int i rdvd real_of_int c * real_of_int x + Inum (real_of_int x # bs) e" 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
  2212
    next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2213
      assume
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2214
        "real_of_int i rdvd real_of_int c * real_of_int x + Inum (real_of_int x # bs) e" (is "?ri rdvd ?rc*?rx+?e")
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2215
      hence "\<exists> (l::int). ?rc*?rx+?e = ?ri * (real_of_int l)" by (simp add: rdvd_def)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2216
      hence "\<exists> (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int l) - real_of_int c * (real_of_int k * real_of_int d)" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2217
      hence "\<exists> (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int l) - real_of_int c * (real_of_int k * real_of_int i * real_of_int di)" by (simp add: di_def)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2218
      hence "\<exists> (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int (l - c*k*di))" by (simp add: algebra_simps)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2219
      hence "\<exists> (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int l)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2220
        by blast
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2221
      thus "real_of_int i rdvd real_of_int c * real_of_int x - real_of_int c * (real_of_int k * real_of_int d) + Inum (real_of_int x # bs) e" 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
  2222
    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
  2223
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
  2224
  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
  2225
    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
  2226
    then obtain "di" where di_def: "d=i*di" by blast
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2227
    show ?case
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2228
    proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="real_of_int x - real_of_int k * real_of_int d" and b'="real_of_int x"] right_diff_distrib, rule iffI)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2229
      assume
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2230
        "real_of_int i rdvd real_of_int c * real_of_int x - real_of_int c * (real_of_int k * real_of_int d) + Inum (real_of_int 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
  2231
      (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt")
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2232
      hence "\<exists> (l::int). ?rt = ?ri * (real_of_int l)" by (simp add: rdvd_def)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2233
      hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real_of_int l)+?rc*(?rk * (real_of_int i) * (real_of_int di))"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2234
        by (simp add: algebra_simps di_def)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2235
      hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real_of_int (l + c*k*di))"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2236
        by (simp add: algebra_simps)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2237
      hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri* (real_of_int l)" by blast
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2238
      thus "real_of_int i rdvd real_of_int c * real_of_int x + Inum (real_of_int x # bs) e" 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
  2239
    next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2240
      assume
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2241
        "real_of_int i rdvd real_of_int c * real_of_int x + Inum (real_of_int x # bs) e" (is "?ri rdvd ?rc*?rx+?e")
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2242
      hence "\<exists> (l::int). ?rc*?rx+?e = ?ri * (real_of_int l)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2243
        by (simp add: rdvd_def)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2244
      hence "\<exists> (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int l) - real_of_int c * (real_of_int k * real_of_int d)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2245
        by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2246
      hence "\<exists> (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int l) - real_of_int c * (real_of_int k * real_of_int i * real_of_int di)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2247
        by (simp add: di_def)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2248
      hence "\<exists> (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int (l - c*k*di))"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2249
        by (simp add: algebra_simps)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2250
      hence "\<exists> (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int l)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2251
        by blast
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2252
      thus "real_of_int i rdvd real_of_int c * real_of_int x - real_of_int c * (real_of_int k * real_of_int d) + Inum (real_of_int x # bs) e"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2253
        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
  2254
    qed
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2255
qed (auto simp add: numbound0_I[where bs="bs" and b="real_of_int(x - k*d)" and b'="real_of_int x"] simp del: of_int_mult 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
  2256
324622260d29 Added 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
lemma minusinf_ex:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2258
  assumes lin: "iszlfm p (real_of_int (a::int) #bs)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2259
  and exmi: "\<exists> (x::int). Ifm (real_of_int x#bs) (minusinf p)" (is "\<exists> x. ?P1 x")
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2260
  shows "\<exists> (x::int). Ifm (real_of_int x#bs) p" (is "\<exists> x. ?P 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
  2261
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
  2262
  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
  2263
  from \<delta> [OF lin] have dpos: "?d >0" by simp
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2264
  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
  2265
  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
  2266
  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
  2267
  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
  2268
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
  2269
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2270
lemma minusinf_bex:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2271
  assumes lin: "iszlfm p (real_of_int (a::int) #bs)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2272
  shows "(\<exists> (x::int). Ifm (real_of_int x#bs) (minusinf p)) =
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2273
         (\<exists> (x::int)\<in> {1..\<delta> p}. Ifm (real_of_int x#bs) (minusinf 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
  2274
  (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
  2275
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
  2276
  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
  2277
  from \<delta> [OF lin] have dpos: "?d >0" by simp
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2278
  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
  2279
  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
  2280
  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
  2281
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
  2282
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2283
lemma dvd1_eq1: "x > 0 \<Longrightarrow> is_unit x \<longleftrightarrow> x = 1" for x :: int
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2284
  by simp
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2285
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2286
fun a_\<beta> :: "fm \<Rightarrow> int \<Rightarrow> fm" (* adjusts the coefficients of a formula *)
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2287
where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2288
  "a_\<beta> (And p q) = (\<lambda> k. And (a_\<beta> p k) (a_\<beta> q k))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2289
| "a_\<beta> (Or p q) = (\<lambda> k. Or (a_\<beta> p k) (a_\<beta> q k))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2290
| "a_\<beta> (Eq  (CN 0 c e)) = (\<lambda> k. Eq (CN 0 1 (Mul (k div c) e)))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2291
| "a_\<beta> (NEq (CN 0 c e)) = (\<lambda> k. NEq (CN 0 1 (Mul (k div c) e)))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2292
| "a_\<beta> (Lt  (CN 0 c e)) = (\<lambda> k. Lt (CN 0 1 (Mul (k div c) e)))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2293
| "a_\<beta> (Le  (CN 0 c e)) = (\<lambda> k. Le (CN 0 1 (Mul (k div c) e)))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2294
| "a_\<beta> (Gt  (CN 0 c e)) = (\<lambda> k. Gt (CN 0 1 (Mul (k div c) e)))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2295
| "a_\<beta> (Ge  (CN 0 c e)) = (\<lambda> k. Ge (CN 0 1 (Mul (k div c) e)))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2296
| "a_\<beta> (Dvd i (CN 0 c e)) =(\<lambda> k. Dvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2297
| "a_\<beta> (NDvd i (CN 0 c e))=(\<lambda> k. NDvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2298
| "a_\<beta> p = (\<lambda> k. p)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2299
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2300
fun d_\<beta> :: "fm \<Rightarrow> int \<Rightarrow> bool" (* tests if all coeffs c of c divide a given l*)
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2301
where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2302
  "d_\<beta> (And p q) = (\<lambda> k. (d_\<beta> p k) \<and> (d_\<beta> q k))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2303
| "d_\<beta> (Or p q) = (\<lambda> k. (d_\<beta> p k) \<and> (d_\<beta> q k))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2304
| "d_\<beta> (Eq  (CN 0 c e)) = (\<lambda> k. c dvd k)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2305
| "d_\<beta> (NEq (CN 0 c e)) = (\<lambda> k. c dvd k)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2306
| "d_\<beta> (Lt  (CN 0 c e)) = (\<lambda> k. c dvd k)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2307
| "d_\<beta> (Le  (CN 0 c e)) = (\<lambda> k. c dvd k)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2308
| "d_\<beta> (Gt  (CN 0 c e)) = (\<lambda> k. c dvd k)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2309
| "d_\<beta> (Ge  (CN 0 c e)) = (\<lambda> k. c dvd k)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2310
| "d_\<beta> (Dvd i (CN 0 c e)) =(\<lambda> k. c dvd k)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2311
| "d_\<beta> (NDvd i (CN 0 c e))=(\<lambda> k. c dvd k)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2312
| "d_\<beta> p = (\<lambda> k. True)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2313
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2314
fun \<zeta>  :: "fm \<Rightarrow> int" (* computes the lcm of all coefficients of x*)
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2315
where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2316
  "\<zeta> (And p q) = lcm (\<zeta> p) (\<zeta> q)"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2317
| "\<zeta> (Or p q) = lcm (\<zeta> p) (\<zeta> q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2318
| "\<zeta> (Eq  (CN 0 c e)) = c"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2319
| "\<zeta> (NEq (CN 0 c e)) = c"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2320
| "\<zeta> (Lt  (CN 0 c e)) = c"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2321
| "\<zeta> (Le  (CN 0 c e)) = c"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2322
| "\<zeta> (Gt  (CN 0 c e)) = c"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2323
| "\<zeta> (Ge  (CN 0 c e)) = c"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2324
| "\<zeta> (Dvd i (CN 0 c e)) = c"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2325
| "\<zeta> (NDvd i (CN 0 c e))= c"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2326
| "\<zeta> p = 1"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2327
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2328
fun \<beta> :: "fm \<Rightarrow> num list"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2329
where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2330
  "\<beta> (And p q) = (\<beta> p @ \<beta> q)"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2331
| "\<beta> (Or p q) = (\<beta> p @ \<beta> q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2332
| "\<beta> (Eq  (CN 0 c e)) = [Sub (C (- 1)) e]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2333
| "\<beta> (NEq (CN 0 c e)) = [Neg e]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2334
| "\<beta> (Lt  (CN 0 c e)) = []"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2335
| "\<beta> (Le  (CN 0 c e)) = []"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2336
| "\<beta> (Gt  (CN 0 c e)) = [Neg e]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2337
| "\<beta> (Ge  (CN 0 c e)) = [Sub (C (- 1)) e]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2338
| "\<beta> p = []"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2339
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2340
fun \<alpha> :: "fm \<Rightarrow> num list"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2341
where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2342
  "\<alpha> (And p q) = (\<alpha> p @ \<alpha> q)"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2343
| "\<alpha> (Or p q) = (\<alpha> p @ \<alpha> q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2344
| "\<alpha> (Eq  (CN 0 c e)) = [Add (C (- 1)) e]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2345
| "\<alpha> (NEq (CN 0 c e)) = [e]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2346
| "\<alpha> (Lt  (CN 0 c e)) = [e]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2347
| "\<alpha> (Le  (CN 0 c e)) = [Add (C (- 1)) e]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2348
| "\<alpha> (Gt  (CN 0 c e)) = []"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2349
| "\<alpha> (Ge  (CN 0 c e)) = []"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2350
| "\<alpha> p = []"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2351
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2352
fun mirror :: "fm \<Rightarrow> fm"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2353
where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2354
  "mirror (And p q) = And (mirror p) (mirror q)"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2355
| "mirror (Or p q) = Or (mirror p) (mirror q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2356
| "mirror (Eq  (CN 0 c e)) = Eq (CN 0 c (Neg e))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2357
| "mirror (NEq (CN 0 c e)) = NEq (CN 0 c (Neg e))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2358
| "mirror (Lt  (CN 0 c e)) = Gt (CN 0 c (Neg e))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2359
| "mirror (Le  (CN 0 c e)) = Ge (CN 0 c (Neg e))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2360
| "mirror (Gt  (CN 0 c e)) = Lt (CN 0 c (Neg e))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2361
| "mirror (Ge  (CN 0 c e)) = Le (CN 0 c (Neg e))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2362
| "mirror (Dvd i (CN 0 c e)) = Dvd i (CN 0 c (Neg e))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2363
| "mirror (NDvd i (CN 0 c e)) = NDvd i (CN 0 c (Neg e))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2364
| "mirror 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
  2365
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2366
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
  2367
  assumes lp: "iszlfm p (a#bs)"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2368
  shows "(Inum (real_of_int (i::int)#bs)) ` set (\<alpha> p) = (Inum (real_of_int i#bs)) ` set (\<beta> (mirror p))"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2369
  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
  2370
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2371
lemma mirror:
23264
324622260d29 Added 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
  assumes lp: "iszlfm p (a#bs)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2373
  shows "Ifm (real_of_int (x::int)#bs) (mirror p) = Ifm (real_of_int (- x)#bs) p"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2374
  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
  2375
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
  2376
  case (9 j c e)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2377
  have th: "(real_of_int j rdvd real_of_int c * real_of_int x - Inum (real_of_int x # bs) e) =
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2378
       (real_of_int j rdvd - (real_of_int c * real_of_int x - Inum (real_of_int 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
  2379
    by (simp only: rdvd_minus[symmetric])
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2380
  from 9 th show ?case
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  2381
    by (simp add: algebra_simps
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2382
      numbound0_I[where bs="bs" and b'="real_of_int x" and b="- real_of_int 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
  2383
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2384
  case (10 j c e)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2385
  have th: "(real_of_int j rdvd real_of_int c * real_of_int x - Inum (real_of_int x # bs) e) =
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2386
       (real_of_int j rdvd - (real_of_int c * real_of_int x - Inum (real_of_int 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
  2387
    by (simp only: rdvd_minus[symmetric])
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2388
  from 10 th show  ?case
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  2389
    by (simp add: algebra_simps
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2390
      numbound0_I[where bs="bs" and b'="real_of_int x" and b="- real_of_int x"])
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2391
qed (auto simp add: numbound0_I[where bs="bs" and b="real_of_int x" and b'="- real_of_int 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
  2392
324622260d29 Added 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
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
  2394
  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
  2395
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2396
lemma mirror_d_\<beta>: "iszlfm p (a#bs) \<and> d_\<beta> p 1
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2397
  \<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
  2398
  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
  2399
324622260d29 Added 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
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
  2401
  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
  2402
324622260d29 Added 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
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2404
lemma mirror_ex:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2405
  assumes lp: "iszlfm p (real_of_int (i::int)#bs)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2406
  shows "(\<exists> (x::int). Ifm (real_of_int x#bs) (mirror p)) = (\<exists> (x::int). Ifm (real_of_int 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
  2407
  (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
  2408
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
  2409
  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
  2410
  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
  2411
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2412
  fix x assume "?I x p" hence "?I (- x) ?mp"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2413
    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
  2414
  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
  2415
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
  2416
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2417
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
  2418
  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
  2419
  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
  2420
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2421
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
  2422
  assumes linp: "iszlfm p (a #bs)"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2423
  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
  2424
  and d: "l dvd l'"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2425
  shows "d_\<beta> p l'"
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
  2426
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
  2427
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
  2428
324622260d29 Added 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
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
  2430
  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
  2431
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
  2432
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
  2433
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2434
lemma \<zeta>:
23264
324622260d29 Added 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
  assumes linp: "iszlfm p (a #bs)"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2436
  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
  2437
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
  2438
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
  2439
  case (1 p q)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2440
  then  have dl1: "\<zeta> p dvd lcm (\<zeta> p) (\<zeta> q)" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2441
  from 1 have dl2: "\<zeta> q dvd lcm (\<zeta> p) (\<zeta> q)" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2442
  from 1 d_\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="lcm (\<zeta> p) (\<zeta> q)"]
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2443
    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
  2444
    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
  2445
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
  2446
  case (2 p q)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2447
  then have dl1: "\<zeta> p dvd lcm (\<zeta> p) (\<zeta> q)" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2448
  from 2 have dl2: "\<zeta> q dvd lcm (\<zeta> p) (\<zeta> q)" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2449
  from 2 d_\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="lcm (\<zeta> p) (\<zeta> q)"]
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2450
    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
  2451
    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
  2452
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
  2453
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2454
lemma a_\<beta>: assumes linp: "iszlfm p (a #bs)" and d: "d_\<beta> p l" and lp: "l > 0"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2455
  shows "iszlfm (a_\<beta> p l) (a #bs) \<and> d_\<beta> (a_\<beta> p l) 1 \<and> (Ifm (real_of_int (l * x) #bs) (a_\<beta> p l) = Ifm ((real_of_int 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
  2456
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
  2457
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
  2458
  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
  2459
    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
  2460
    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
  2461
    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
  2462
      by (simp add: zdiv_mono1[OF clel cp])
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2463
    then have ldcp:"0 < l div c"
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  2464
      by (simp add: div_self[OF cnz])
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
  2465
    have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
64246
15d1ee6e847b eliminated irregular aliasses
haftmann
parents: 64240
diff changeset
  2466
    hence cl:"c * (l div c) =l" using mult_div_mod_eq [where a="l" and b="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
  2467
      by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2468
    hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e < (0::real)) =
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2469
          (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e < 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
  2470
      by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2471
    also have "\<dots> = (real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e) < (real_of_int (l div c)) * 0)" by (simp add: algebra_simps)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2472
    also have "\<dots> = (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e < 0)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2473
    using mult_less_0_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e"] ldcp by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2474
  finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"] be  isint_Mul[OF ei] 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
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
  2476
  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
  2477
    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
  2478
    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
  2479
    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
  2480
      by (simp add: zdiv_mono1[OF clel cp])
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2481
    then have ldcp:"0 < l div c"
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  2482
      by (simp add: div_self[OF cnz])
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
  2483
    have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
64246
15d1ee6e847b eliminated irregular aliasses
haftmann
parents: 64240
diff changeset
  2484
    hence cl:"c * (l div c) =l" using mult_div_mod_eq [where a="l" and b="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
  2485
      by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2486
    hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e \<le> (0::real)) =
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2487
          (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e \<le> 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
  2488
      by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2489
    also have "\<dots> = (real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e) \<le> (real_of_int (l div c)) * 0)" by (simp add: algebra_simps)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2490
    also have "\<dots> = (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \<le> 0)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2491
    using mult_le_0_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e"] ldcp by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2492
  finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"]  be  isint_Mul[OF ei] 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
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
  2494
  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
  2495
    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
  2496
    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
  2497
    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
  2498
      by (simp add: zdiv_mono1[OF clel cp])
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2499
    then have ldcp:"0 < l div c"
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  2500
      by (simp add: div_self[OF cnz])
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
  2501
    have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
64246
15d1ee6e847b eliminated irregular aliasses
haftmann
parents: 64240
diff changeset
  2502
    hence cl:"c * (l div c) =l" using mult_div_mod_eq [where a="l" and b="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
  2503
      by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2504
    hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e > (0::real)) =
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2505
          (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e > 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
  2506
      by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2507
    also have "\<dots> = (real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e) > (real_of_int (l div c)) * 0)" by (simp add: algebra_simps)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2508
    also have "\<dots> = (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e > 0)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2509
    using zero_less_mult_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e"] ldcp by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2510
  finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"]  be  isint_Mul[OF ei] 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
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
  2512
  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
  2513
    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
  2514
    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
  2515
    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
  2516
      by (simp add: zdiv_mono1[OF clel cp])
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2517
    then have ldcp:"0 < l div c"
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  2518
      by (simp add: div_self[OF cnz])
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
  2519
    have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
64246
15d1ee6e847b eliminated irregular aliasses
haftmann
parents: 64240
diff changeset
  2520
    hence cl:"c * (l div c) =l" using mult_div_mod_eq [where a="l" and b="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
  2521
      by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2522
    hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e \<ge> (0::real)) =
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2523
          (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e \<ge> 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
  2524
      by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2525
    also have "\<dots> = (real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e) \<ge> (real_of_int (l div c)) * 0)" by (simp add: algebra_simps)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2526
    also have "\<dots> = (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \<ge> 0)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2527
    using zero_le_mult_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e"] ldcp by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2528
  finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"]  be  isint_Mul[OF ei] 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
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
  2530
  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
  2531
    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
  2532
    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
  2533
    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
  2534
      by (simp add: zdiv_mono1[OF clel cp])
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2535
    then have ldcp:"0 < l div c"
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  2536
      by (simp add: div_self[OF cnz])
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
  2537
    have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
64246
15d1ee6e847b eliminated irregular aliasses
haftmann
parents: 64240
diff changeset
  2538
    hence cl:"c * (l div c) =l" using mult_div_mod_eq [where a="l" and b="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
  2539
      by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2540
    hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = (0::real)) =
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2541
          (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = 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
  2542
      by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2543
    also have "\<dots> = (real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e) = (real_of_int (l div c)) * 0)" by (simp add: algebra_simps)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2544
    also have "\<dots> = (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e = 0)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2545
    using mult_eq_0_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e"] ldcp by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2546
  finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"]  be  isint_Mul[OF ei] 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
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
  2548
  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
  2549
    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
  2550
    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
  2551
    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
  2552
      by (simp add: zdiv_mono1[OF clel cp])
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2553
    then have ldcp:"0 < l div c"
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  2554
      by (simp add: div_self[OF cnz])
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
  2555
    have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
64246
15d1ee6e847b eliminated irregular aliasses
haftmann
parents: 64240
diff changeset
  2556
    hence cl:"c * (l div c) =l" using mult_div_mod_eq [where a="l" and b="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
  2557
      by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2558
    hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e \<noteq> (0::real)) =
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2559
          (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) 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
  2560
      by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2561
    also have "\<dots> = (real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e) \<noteq> (real_of_int (l div c)) * 0)" by (simp add: algebra_simps)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2562
    also have "\<dots> = (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \<noteq> 0)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2563
    using zero_le_mult_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e"] ldcp by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2564
  finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"]  be  isint_Mul[OF ei] 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
  2565
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
  2566
  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
  2567
    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
  2568
    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
  2569
    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
  2570
      by (simp add: zdiv_mono1[OF clel cp])
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2571
    then have ldcp:"0 < l div c"
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  2572
      by (simp add: div_self[OF cnz])
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
  2573
    have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
64246
15d1ee6e847b eliminated irregular aliasses
haftmann
parents: 64240
diff changeset
  2574
    hence cl:"c * (l div c) =l" using mult_div_mod_eq [where a="l" and b="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
  2575
      by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2576
    hence "(\<exists> (k::int). real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = (real_of_int (l div c) * real_of_int j) * real_of_int k) = (\<exists> (k::int). real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = (real_of_int (l div c) * real_of_int j) * real_of_int k)"  by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2577
    also have "\<dots> = (\<exists> (k::int). real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e - real_of_int j * real_of_int k) = real_of_int (l div c)*0)" by (simp add: algebra_simps)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2578
    also fix k have "\<dots> = (\<exists> (k::int). real_of_int c * real_of_int x + Inum (real_of_int x # bs) e - real_of_int j * real_of_int k = 0)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2579
    using zero_le_mult_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e - real_of_int j * real_of_int k"] ldcp by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2580
  also have "\<dots> = (\<exists> (k::int). real_of_int c * real_of_int x + Inum (real_of_int x # bs) e = real_of_int j * real_of_int k)" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2581
  finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"] rdvd_def  be  isint_Mul[OF ei] mult_strict_mono[OF ldcp jp ldcp ] 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
  2582
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
  2583
  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
  2584
    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
  2585
    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
  2586
    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
  2587
      by (simp add: zdiv_mono1[OF clel cp])
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2588
    then have ldcp:"0 < l div c"
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  2589
      by (simp add: div_self[OF cnz])
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
  2590
    have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
64246
15d1ee6e847b eliminated irregular aliasses
haftmann
parents: 64240
diff changeset
  2591
    hence cl:"c * (l div c) =l" using mult_div_mod_eq [where a="l" and b="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
  2592
      by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2593
    hence "(\<exists> (k::int). real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = (real_of_int (l div c) * real_of_int j) * real_of_int k) = (\<exists> (k::int). real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = (real_of_int (l div c) * real_of_int j) * real_of_int k)"  by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2594
    also have "\<dots> = (\<exists> (k::int). real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e - real_of_int j * real_of_int k) = real_of_int (l div c)*0)" by (simp add: algebra_simps)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2595
    also fix k have "\<dots> = (\<exists> (k::int). real_of_int c * real_of_int x + Inum (real_of_int x # bs) e - real_of_int j * real_of_int k = 0)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2596
    using zero_le_mult_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e - real_of_int j * real_of_int k"] ldcp by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2597
  also have "\<dots> = (\<exists> (k::int). real_of_int c * real_of_int x + Inum (real_of_int x # bs) e = real_of_int j * real_of_int k)" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2598
  finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"] rdvd_def  be  isint_Mul[OF ei]  mult_strict_mono[OF ldcp jp ldcp ] by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2599
qed (simp_all add: numbound0_I[where bs="bs" and b="real_of_int (l * x)" and b'="real_of_int x"] isint_Mul del: 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
  2600
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2601
lemma a_\<beta>_ex: assumes linp: "iszlfm p (a#bs)" and d: "d_\<beta> p l" and lp: "l>0"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2602
  shows "(\<exists> x. l dvd x \<and> Ifm (real_of_int x #bs) (a_\<beta> p l)) = (\<exists> (x::int). Ifm (real_of_int 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
  2603
  (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
  2604
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
  2605
  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
  2606
    using unity_coeff_ex[where l="l" and P="?P", simplified] by simp
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2607
  also have "\<dots> = (\<exists> (x::int). ?P' x)" using a_\<beta>[OF linp d lp] by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2608
  finally show ?thesis  .
23264
324622260d29 Added 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
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
  2610
324622260d29 Added 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
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
  2612
  assumes lp: "iszlfm p (a#bs)"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2613
  and u: "d_\<beta> p 1"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2614
  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
  2615
  and dp: "d > 0"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2616
  and nob: "\<not>(\<exists>(j::int) \<in> {1 .. d}. \<exists> b\<in> (Inum (a#bs)) ` set(\<beta> p). real_of_int x = b + real_of_int j)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2617
  and p: "Ifm (real_of_int x#bs) p" (is "?P 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
  2618
  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
  2619
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
  2620
proof(induct p rule: iszlfm.induct)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2621
  case (5 c e) hence c1: "c=1" and  bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp_all
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2622
  with dp p c1 numbound0_I[OF bn,where b="real_of_int (x-d)" and b'="real_of_int x" and bs="bs"] 5
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2623
  show ?case by (simp del: 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
  2624
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2625
  case (6 c e)  hence c1: "c=1" and  bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp_all
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2626
  with dp p c1 numbound0_I[OF bn,where b="real_of_int (x-d)" and b'="real_of_int x" and bs="bs"] 6
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2627
  show ?case by (simp del: 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
  2628
next
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2629
  case (7 c e) hence p: "Ifm (real_of_int x #bs) (Gt (CN 0 c e))" and c1: "c=1"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2630
    and bn:"numbound0 e" and ie1:"isint e (a#bs)" using dvd1_eq1[where x="c"] by simp_all
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2631
  let ?e = "Inum (real_of_int x # bs) e"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2632
  from ie1 have ie: "real_of_int \<lfloor>?e\<rfloor> = ?e" using isint_iff[where n="e" and bs="a#bs"]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2633
      numbound0_I[OF bn,where b="a" and b'="real_of_int x" and bs="bs"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2634
    by (simp add: isint_iff)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2635
    {assume "real_of_int (x-d) +?e > 0" hence ?case using c1
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2636
      numbound0_I[OF bn,where b="real_of_int (x-d)" and b'="real_of_int x" and bs="bs"]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2637
        by (simp del: 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
  2638
    moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2639
    {assume H: "\<not> real_of_int (x-d) + ?e > 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
  2640
      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
  2641
      have vb: "?v \<in> set (\<beta> (Gt (CN 0 c e)))" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2642
      from 7(5)[simplified simp_thms Inum.simps \<beta>.simps list.set bex_simps numbound0_I[OF bn,where b="a" and b'="real_of_int x" and bs="bs"]]
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2643
      have nob: "\<not> (\<exists> j\<in> {1 ..d}. real_of_int x =  - ?e + real_of_int j)" by auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2644
      from H p have "real_of_int x + ?e > 0 \<and> real_of_int x + ?e \<le> real_of_int d" by (simp add: c1)
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2645
      hence "real_of_int (x + \<lfloor>?e\<rfloor>) > real_of_int (0::int) \<and> real_of_int (x + \<lfloor>?e\<rfloor>) \<le> real_of_int d"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2646
        using ie by simp
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2647
      hence "x + \<lfloor>?e\<rfloor> \<ge> 1 \<and> x + \<lfloor>?e\<rfloor> \<le> d"  by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2648
      hence "\<exists> (j::int) \<in> {1 .. d}. j = x + \<lfloor>?e\<rfloor>" by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2649
      hence "\<exists> (j::int) \<in> {1 .. d}. real_of_int x = real_of_int (- \<lfloor>?e\<rfloor> + j)" by force
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2650
      hence "\<exists> (j::int) \<in> {1 .. d}. real_of_int x = - ?e + real_of_int j"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2651
        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
  2652
      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
  2653
    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
  2654
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2655
  case (8 c e) hence p: "Ifm (real_of_int x #bs) (Ge (CN 0 c e))" and c1: "c=1" and bn:"numbound0 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
  2656
    and ie1:"isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2657
    let ?e = "Inum (real_of_int x # bs) e"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2658
    from ie1 have ie: "real_of_int \<lfloor>?e\<rfloor> = ?e" using numbound0_I[OF bn,where b="real_of_int x" and b'="a" and bs="bs"] isint_iff[where n="e" and bs="(real_of_int x)#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
  2659
      by (simp add: isint_iff)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2660
    {assume "real_of_int (x-d) +?e \<ge> 0" hence ?case using  c1
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2661
      numbound0_I[OF bn,where b="real_of_int (x-d)" and b'="real_of_int x" and bs="bs"]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2662
        by (simp del: 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
  2663
    moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2664
    {assume H: "\<not> real_of_int (x-d) + ?e \<ge> 0"
58410
6d46ad54a2ab explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents: 58310
diff changeset
  2665
      let ?v="Sub (C (- 1)) 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
  2666
      have vb: "?v \<in> set (\<beta> (Ge (CN 0 c e)))" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2667
      from 8(5)[simplified simp_thms Inum.simps \<beta>.simps list.set bex_simps numbound0_I[OF bn,where b="a" and b'="real_of_int x" and bs="bs"]]
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2668
      have nob: "\<not> (\<exists> j\<in> {1 ..d}. real_of_int x =  - ?e - 1 + real_of_int j)" by auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2669
      from H p have "real_of_int x + ?e \<ge> 0 \<and> real_of_int x + ?e < real_of_int d" by (simp add: c1)
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2670
      hence "real_of_int (x + \<lfloor>?e\<rfloor>) \<ge> real_of_int (0::int) \<and> real_of_int (x + \<lfloor>?e\<rfloor>) < real_of_int d"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2671
        using ie by simp
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2672
      hence "x + \<lfloor>?e\<rfloor> + 1 \<ge> 1 \<and> x + \<lfloor>?e\<rfloor> + 1 \<le> d" by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2673
      hence "\<exists> (j::int) \<in> {1 .. d}. j = x + \<lfloor>?e\<rfloor> + 1" by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2674
      hence "\<exists> (j::int) \<in> {1 .. d}. x= - \<lfloor>?e\<rfloor> - 1 + j" by (simp add: algebra_simps)
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2675
      hence "\<exists> (j::int) \<in> {1 .. d}. real_of_int x= real_of_int (- \<lfloor>?e\<rfloor> - 1 + j)" by presburger
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2676
      hence "\<exists> (j::int) \<in> {1 .. d}. real_of_int x= - ?e - 1 + real_of_int j"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2677
        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
  2678
      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
  2679
    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
  2680
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2681
  case (3 c e) hence p: "Ifm (real_of_int x #bs) (Eq (CN 0 c e))" (is "?p x") and c1: "c=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
  2682
    and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2683
    let ?e = "Inum (real_of_int x # bs) e"
58410
6d46ad54a2ab explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents: 58310
diff changeset
  2684
    let ?v="(Sub (C (- 1)) 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
  2685
    have vb: "?v \<in> set (\<beta> (Eq (CN 0 c e)))" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2686
    from p have "real_of_int 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
  2687
      by simp (erule ballE[where x="1"],
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2688
        simp_all add:algebra_simps numbound0_I[OF bn,where b="real_of_int 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
  2689
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2690
  case (4 c e)hence p: "Ifm (real_of_int x #bs) (NEq (CN 0 c e))" (is "?p x") and c1: "c=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
  2691
    and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2692
    let ?e = "Inum (real_of_int 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
  2693
    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
  2694
    have vb: "?v \<in> set (\<beta> (NEq (CN 0 c e)))" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2695
    {assume "real_of_int x - real_of_int d + Inum ((real_of_int (x -d)) # bs) 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
  2696
      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
  2697
    moreover
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2698
    {assume H: "real_of_int x - real_of_int d + Inum ((real_of_int (x -d)) # bs) e = 0"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2699
      hence "real_of_int x = - Inum ((real_of_int (x -d)) # bs) e + real_of_int d" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2700
      hence "real_of_int x = - Inum (a # bs) e + real_of_int d"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2701
        by (simp add: numbound0_I[OF bn,where b="real_of_int x - real_of_int d"and b'="a"and bs="bs"])
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2702
       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
  2703
  ultimately show ?case by blast
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2704
next
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2705
  case (9 j c e) hence p: "Ifm (real_of_int x #bs) (Dvd j (CN 0 c e))" (is "?p x") and c1: "c=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
  2706
    and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2707
  let ?e = "Inum (real_of_int x # bs) e"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2708
  from 9 have "isint e (a #bs)"  by simp
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2709
  hence ie: "real_of_int \<lfloor>?e\<rfloor> = ?e" using isint_iff[where n="e" and bs="(real_of_int x)#bs"] numbound0_I[OF bn,where b="real_of_int x" and b'="a" and bs="bs"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2710
    by (simp add: isint_iff)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2711
  from 9 have id: "j dvd d" by simp
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2712
  from c1 ie[symmetric] have "?p x = (real_of_int j rdvd real_of_int (x + \<lfloor>?e\<rfloor>))" by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2713
  also have "\<dots> = (j dvd x + \<lfloor>?e\<rfloor>)"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2714
    using int_rdvd_real[where i="j" and x="real_of_int (x + \<lfloor>?e\<rfloor>)"] by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2715
  also have "\<dots> = (j dvd x - d + \<lfloor>?e\<rfloor>)"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2716
    using dvd_period[OF id, where x="x" and c="-1" and t="\<lfloor>?e\<rfloor>"] by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2717
  also have "\<dots> = (real_of_int j rdvd real_of_int (x - d + \<lfloor>?e\<rfloor>))"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2718
    using int_rdvd_real[where i="j" and x="real_of_int (x - d + \<lfloor>?e\<rfloor>)",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
  2719
      ie by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2720
  also have "\<dots> = (real_of_int j rdvd real_of_int x - real_of_int d + ?e)"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2721
    using ie by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2722
  finally show ?case
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2723
    using numbound0_I[OF bn,where b="real_of_int (x-d)" and b'="real_of_int 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
  2724
next
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2725
  case (10 j c e) hence p: "Ifm (real_of_int 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+
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2726
  let ?e = "Inum (real_of_int x # bs) e"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2727
  from 10 have "isint e (a#bs)"  by simp
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2728
  hence ie: "real_of_int \<lfloor>?e\<rfloor> = ?e" using numbound0_I[OF bn,where b="real_of_int x" and b'="a" and bs="bs"] isint_iff[where n="e" and bs="(real_of_int x)#bs"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2729
    by (simp add: isint_iff)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2730
  from 10 have id: "j dvd d" by simp
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2731
  from c1 ie[symmetric] have "?p x = (\<not> real_of_int j rdvd real_of_int (x + \<lfloor>?e\<rfloor>))" by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2732
  also have "\<dots> = (\<not> j dvd x + \<lfloor>?e\<rfloor>)"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2733
    using int_rdvd_real[where i="j" and x="real_of_int (x + \<lfloor>?e\<rfloor>)"] by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2734
  also have "\<dots> = (\<not> j dvd x - d + \<lfloor>?e\<rfloor>)"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2735
    using dvd_period[OF id, where x="x" and c="-1" and t="\<lfloor>?e\<rfloor>"] by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2736
  also have "\<dots> = (\<not> real_of_int j rdvd real_of_int (x - d + \<lfloor>?e\<rfloor>))"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2737
    using int_rdvd_real[where i="j" and x="real_of_int (x - d + \<lfloor>?e\<rfloor>)",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
  2738
      ie by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2739
  also have "\<dots> = (\<not> real_of_int j rdvd real_of_int x - real_of_int d + ?e)"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2740
    using ie by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2741
  finally show ?case
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2742
    using numbound0_I[OF bn,where b="real_of_int (x-d)" and b'="real_of_int x" and bs="bs"] c1 p by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2743
qed (auto simp add: numbound0_I[where bs="bs" and b="real_of_int (x - d)" and b'="real_of_int x"]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2744
  simp del: 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
  2745
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2746
lemma \<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
  2747
  assumes lp: "iszlfm p (a #bs)"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2748
  and u: "d_\<beta> p 1"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2749
  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
  2750
  and dp: "d > 0"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2751
  shows "\<forall> x. \<not>(\<exists>(j::int) \<in> {1 .. d}. \<exists> b\<in> set(\<beta> p). Ifm ((Inum (a#bs) b + real_of_int j) #bs) p) \<longrightarrow> Ifm (real_of_int x#bs) p \<longrightarrow> Ifm (real_of_int (x - d)#bs) p" (is "\<forall> x. ?b \<longrightarrow> ?P x \<longrightarrow> ?P (x - 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
  2752
proof(clarify)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2753
  fix x
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2754
  assume nb:"?b" and px: "?P x"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2755
  hence nb2: "\<not>(\<exists>(j::int) \<in> {1 .. d}. \<exists> b\<in> (Inum (a#bs)) ` set(\<beta> p). real_of_int x = b + real_of_int 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
  2756
    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
  2757
  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
  2758
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
  2759
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2760
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
  2761
  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
  2762
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
  2763
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  2764
lemma cpmi_eq: "0 < D \<Longrightarrow> (\<exists>z::int. \<forall>x. x < z \<longrightarrow> (P x = P1 x))
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  2765
\<Longrightarrow> \<forall>x. \<not>(\<exists>(j::int) \<in> {1..D}. \<exists>(b::int) \<in> B. P(b+j)) \<longrightarrow> P (x) \<longrightarrow> P (x - D)
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  2766
\<Longrightarrow> (\<forall>(x::int). \<forall>(k::int). ((P1 x)= (P1 (x-k*D))))
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  2767
\<Longrightarrow> (\<exists>(x::int). P(x)) = ((\<exists>(j::int) \<in> {1..D} . (P1(j))) | (\<exists>(j::int) \<in> {1..D}. \<exists>(b::int) \<in> B. P (b+j)))"
23316
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2768
apply(rule iffI)
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2769
prefer 2
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2770
apply(drule minusinfinity)
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2771
apply assumption+
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44121
diff changeset
  2772
apply(fastforce)
23316
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2773
apply clarsimp
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  2774
apply(subgoal_tac "\<And>k. 0<=k \<Longrightarrow> \<forall>x. P x \<longrightarrow> P (x - k*D)")
23316
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2775
apply(frule_tac x = x and z=z in decr_lemma)
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2776
apply(subgoal_tac "P1(x - (\<bar>x - z\<bar> + 1) * D)")
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2777
prefer 2
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2778
apply(subgoal_tac "0 <= (\<bar>x - z\<bar> + 1)")
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2779
prefer 2 apply arith
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44121
diff changeset
  2780
 apply fastforce
23316
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2781
apply(drule (1)  periodic_finite_ex)
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2782
apply blast
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2783
apply(blast dest:decr_mult_lemma)
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2784
done
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2785
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2786
23264
324622260d29 Added 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
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
  2788
  assumes lp: "iszlfm p (a #bs)"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2789
  and u: "d_\<beta> p 1"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2790
  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
  2791
  and dp: "d > 0"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2792
  shows "(\<exists> (x::int). Ifm (real_of_int x #bs) p) = (\<exists> j\<in> {1.. d}. Ifm (real_of_int j #bs) (minusinf p) \<or> (\<exists> b \<in> set (\<beta> p). Ifm ((Inum (a#bs) b + real_of_int j) #bs) p))"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2793
  (is "(\<exists> (x::int). ?P (real_of_int x)) = (\<exists> j\<in> ?D. ?M j \<or> (\<exists> b\<in> ?B. ?P (?I b + real_of_int 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
  2794
proof-
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2795
  from minusinf_inf[OF lp]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2796
  have th: "\<exists>(z::int). \<forall>x<z. ?P (real_of_int x) = ?M x" by blast
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2797
  let ?B' = "{\<lfloor>?I b\<rfloor> | b. b\<in> ?B}"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2798
  from \<beta>_int[OF lp] isint_iff[where bs="a # bs"] have B: "\<forall> b\<in> ?B. real_of_int \<lfloor>?I b\<rfloor> = ?I b" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2799
  from B[rule_format]
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2800
  have "(\<exists>j\<in>?D. \<exists>b\<in> ?B. ?P (?I b + real_of_int j)) = (\<exists>j\<in>?D. \<exists>b\<in> ?B. ?P (real_of_int \<lfloor>?I b\<rfloor> + real_of_int 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
  2801
    by simp
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2802
  also have "\<dots> = (\<exists>j\<in>?D. \<exists>b\<in> ?B. ?P (real_of_int (\<lfloor>?I b\<rfloor> + j)))" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2803
  also have"\<dots> = (\<exists> j \<in> ?D. \<exists> b \<in> ?B'. ?P (real_of_int (b + j)))"  by blast
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2804
  finally have BB':
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2805
    "(\<exists>j\<in>?D. \<exists>b\<in> ?B. ?P (?I b + real_of_int j)) = (\<exists> j \<in> ?D. \<exists> b \<in> ?B'. ?P (real_of_int (b + j)))"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2806
    by blast
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2807
  hence th2: "\<forall> x. \<not> (\<exists> j \<in> ?D. \<exists> b \<in> ?B'. ?P (real_of_int (b + j))) \<longrightarrow> ?P (real_of_int x) \<longrightarrow> ?P (real_of_int (x - d))" using \<beta>'[OF lp u d dp] 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
  2808
  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
  2809
  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
  2810
  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
  2811
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
  2812
324622260d29 Added 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
    (* 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
  2814
324622260d29 Added 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
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2816
fun \<rho> :: "fm \<Rightarrow> (num \<times> int) list" (* Compute the Reddy and Loveland Bset*)
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2817
  where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2818
  "\<rho> (And p q) = (\<rho> p @ \<rho> q)"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2819
| "\<rho> (Or p q) = (\<rho> p @ \<rho> q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2820
| "\<rho> (Eq  (CN 0 c e)) = [(Sub (C (- 1)) e,c)]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2821
| "\<rho> (NEq (CN 0 c e)) = [(Neg e,c)]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2822
| "\<rho> (Lt  (CN 0 c e)) = []"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2823
| "\<rho> (Le  (CN 0 c e)) = []"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2824
| "\<rho> (Gt  (CN 0 c e)) = [(Neg e, c)]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2825
| "\<rho> (Ge  (CN 0 c e)) = [(Sub (C (-1)) e, c)]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2826
| "\<rho> p = []"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2827
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2828
fun \<sigma>_\<rho>:: "fm \<Rightarrow> num \<times> int \<Rightarrow> fm" (* Performs the modified substitution of Reddy and Loveland*)
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2829
where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2830
  "\<sigma>_\<rho> (And p q) = (\<lambda> (t,k). And (\<sigma>_\<rho> p (t,k)) (\<sigma>_\<rho> q (t,k)))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2831
| "\<sigma>_\<rho> (Or p q) = (\<lambda> (t,k). Or (\<sigma>_\<rho> p (t,k)) (\<sigma>_\<rho> q (t,k)))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2832
| "\<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
  2833
                                            else (Eq (Add (Mul c t) (Mul k e))))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2834
| "\<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
  2835
                                            else (NEq (Add (Mul c t) (Mul k e))))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2836
| "\<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
  2837
                                            else (Lt (Add (Mul c t) (Mul k e))))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2838
| "\<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
  2839
                                            else (Le (Add (Mul c t) (Mul k e))))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2840
| "\<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
  2841
                                            else (Gt (Add (Mul c t) (Mul k e))))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2842
| "\<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
  2843
                                            else (Ge (Add (Mul c t) (Mul k e))))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2844
| "\<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
  2845
                                            else (Dvd (i*k) (Add (Mul c t) (Mul k e))))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2846
| "\<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
  2847
                                            else (NDvd (i*k) (Add (Mul c t) (Mul k e))))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2848
| "\<sigma>_\<rho> p = (\<lambda> (t,k). p)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2849
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2850
fun \<alpha>_\<rho> :: "fm \<Rightarrow> (num \<times> int) list"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2851
where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2852
  "\<alpha>_\<rho> (And p q) = (\<alpha>_\<rho> p @ \<alpha>_\<rho> q)"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2853
| "\<alpha>_\<rho> (Or p q) = (\<alpha>_\<rho> p @ \<alpha>_\<rho> q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2854
| "\<alpha>_\<rho> (Eq  (CN 0 c e)) = [(Add (C (- 1)) e,c)]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2855
| "\<alpha>_\<rho> (NEq (CN 0 c e)) = [(e,c)]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2856
| "\<alpha>_\<rho> (Lt  (CN 0 c e)) = [(e,c)]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2857
| "\<alpha>_\<rho> (Le  (CN 0 c e)) = [(Add (C (- 1)) e,c)]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  2858
| "\<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
  2859
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2860
    (* 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
  2861
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  2862
definition \<sigma> :: "fm \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm" where
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2863
  "\<sigma> p k t \<equiv> And (Dvd k t) (\<sigma>_\<rho> p (t,k))"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2864
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2865
lemma \<sigma>_\<rho>:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2866
  assumes linp: "iszlfm p (real_of_int (x::int)#bs)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2867
  and kpos: "real_of_int 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
  2868
  and tnb: "numbound0 t"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2869
  and tint: "isint t (real_of_int x#bs)"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2870
  and kdt: "k dvd \<lfloor>Inum (b'#bs) t\<rfloor>"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2871
  shows "Ifm (real_of_int x#bs) (\<sigma>_\<rho> p (t,k)) = (Ifm ((real_of_int (\<lfloor>Inum (b'#bs) t\<rfloor> div k))#bs) p)"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2872
  (is "?I (real_of_int x) (?s p) = (?I (real_of_int (\<lfloor>?N b' t\<rfloor> div k)) p)" is "_ = (?I ?tk 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
  2873
using linp kpos tnb
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2874
proof(induct p rule: \<sigma>_\<rho>.induct)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2875
  case (3 c e)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2876
  from 3 have cp: "c > 0" and nb: "numbound0 e" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2877
  { assume kdc: "k dvd c"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2878
    from tint have ti: "real_of_int \<lfloor>?N (real_of_int x) t\<rfloor> = ?N (real_of_int x) t" using isint_def by simp
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2879
    from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2880
      numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2881
      numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) }
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2882
  moreover
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2883
  { assume *: "\<not> k dvd c"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2884
    from kpos have knz': "real_of_int k \<noteq> 0" by simp
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2885
    from tint have ti: "real_of_int \<lfloor>?N (real_of_int x) t\<rfloor> = ?N (real_of_int x) t"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2886
      using isint_def by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2887
    from assms * have "?I (real_of_int x) (?s (Eq (CN 0 c e))) = ((real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k = 0)"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2888
      using real_of_int_div[OF kdt]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2889
        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2890
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2891
      by (simp add: ti algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2892
      also have "\<dots> = (?I ?tk (Eq (CN 0 c e)))"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2893
        using nonzero_eq_divide_eq[OF knz',
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2894
            where a="real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e" and b="0", symmetric]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2895
          real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2896
          numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"]
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2897
        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
  2898
      finally have ?case . }
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2899
    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
  2900
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2901
  case (4 c e)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2902
  then have cp: "c > 0" and nb: "numbound0 e" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2903
  { assume kdc: "k dvd c"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2904
    from tint have ti: "real_of_int \<lfloor>?N (real_of_int x) t\<rfloor> = ?N (real_of_int x) t" using isint_def by simp
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2905
    from kdc have  ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2906
      numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2907
      numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) }
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2908
  moreover
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2909
  { assume *: "\<not> k dvd c"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2910
    from kpos have knz': "real_of_int k \<noteq> 0" by simp
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2911
    from tint have ti: "real_of_int \<lfloor>?N (real_of_int x) t\<rfloor> = ?N (real_of_int x) t" using isint_def by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2912
    from assms * have "?I (real_of_int x) (?s (NEq (CN 0 c e))) = ((real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k \<noteq> 0)"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2913
      using real_of_int_div[OF kdt]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2914
        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2915
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2916
      by (simp add: ti algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2917
    also have "\<dots> = (?I ?tk (NEq (CN 0 c e)))"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2918
      using nonzero_eq_divide_eq[OF knz',
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2919
          where a="real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e" and b="0", symmetric]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2920
        real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2921
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2922
      by (simp add: ti)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2923
    finally have ?case . }
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2924
  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
  2925
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2926
  case (5 c e)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2927
  then have cp: "c > 0" and nb: "numbound0 e" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2928
  { assume kdc: "k dvd c"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2929
    from tint have ti: "real_of_int \<lfloor>?N (real_of_int x) t\<rfloor> = ?N (real_of_int x) t" using isint_def by simp
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2930
    from kdc have  ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2931
      numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2932
      numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) }
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2933
  moreover
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2934
  { assume *: "\<not> k dvd c"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2935
    from tint have ti: "real_of_int \<lfloor>?N (real_of_int x) t\<rfloor> = ?N (real_of_int x) t" using isint_def by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2936
    from assms * have "?I (real_of_int x) (?s (Lt (CN 0 c e))) = ((real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k < 0)"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2937
      using real_of_int_div[OF kdt]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2938
        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2939
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2940
      by (simp add: ti algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2941
    also have "\<dots> = (?I ?tk (Lt (CN 0 c e)))"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2942
      using pos_less_divide_eq[OF kpos,
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2943
          where a="real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e" and b="0", symmetric]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2944
        real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2945
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2946
      by (simp add: ti)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2947
    finally have ?case . }
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2948
  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
  2949
next
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2950
  case (6 c e)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2951
  then have cp: "c > 0" and nb: "numbound0 e" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2952
  { assume kdc: "k dvd c"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2953
    from tint have ti: "real_of_int \<lfloor>?N (real_of_int x) t\<rfloor> = ?N (real_of_int x) t" using isint_def by simp
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2954
    from kdc have  ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2955
      numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2956
      numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) }
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2957
  moreover
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2958
  { assume *: "\<not> k dvd c"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2959
    from tint have ti: "real_of_int \<lfloor>?N (real_of_int x) t\<rfloor> = ?N (real_of_int x) t" using isint_def by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2960
    from assms * have "?I (real_of_int x) (?s (Le (CN 0 c e))) = ((real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k \<le> 0)"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2961
      using real_of_int_div[OF kdt]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2962
        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2963
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2964
      by (simp add: ti algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2965
    also have "\<dots> = (?I ?tk (Le (CN 0 c e)))"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2966
      using pos_le_divide_eq[OF kpos,
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2967
          where a="real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e" and b="0", symmetric]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2968
        real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2969
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2970
      by (simp add: ti)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2971
    finally have ?case . }
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2972
  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
  2973
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2974
  case (7 c e)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2975
  then have cp: "c > 0" and nb: "numbound0 e" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2976
  { assume kdc: "k dvd c"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2977
    from tint have ti: "real_of_int \<lfloor>?N (real_of_int x) t\<rfloor> = ?N (real_of_int x) t" using isint_def by simp
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2978
    from kdc have  ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2979
      numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2980
      numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) }
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2981
  moreover
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2982
  { assume *: "\<not> k dvd c"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  2983
    from tint have ti: "real_of_int \<lfloor>?N (real_of_int x) t\<rfloor> = ?N (real_of_int x) t" using isint_def by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2984
    from assms * have "?I (real_of_int x) (?s (Gt (CN 0 c e))) = ((real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k > 0)"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2985
      using real_of_int_div[OF kdt]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2986
        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2987
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2988
      by (simp add: ti algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2989
    also have "\<dots> = (?I ?tk (Gt (CN 0 c e)))"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2990
      using pos_divide_less_eq[OF kpos,
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2991
          where a="real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e" and b="0", symmetric]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2992
        real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2993
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2994
      by (simp add: ti)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2995
    finally have ?case . }
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2996
  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
  2997
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  2998
  case (8 c e)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2999
  then have cp: "c > 0" and nb: "numbound0 e" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3000
  { assume kdc: "k dvd c"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3001
    from tint have ti: "real_of_int \<lfloor>?N (real_of_int x) t\<rfloor> = ?N (real_of_int x) t" using isint_def by simp
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  3002
    from kdc have  ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3003
      numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3004
      numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) }
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3005
  moreover
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3006
  { assume *: "\<not> k dvd c"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3007
    from tint have ti: "real_of_int \<lfloor>?N (real_of_int x) t\<rfloor> = ?N (real_of_int x) t" using isint_def by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3008
    from assms * have "?I (real_of_int x) (?s (Ge (CN 0 c e))) = ((real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k \<ge> 0)"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  3009
      using real_of_int_div[OF kdt]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3010
        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3011
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3012
      by (simp add: ti algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3013
    also have "\<dots> = (?I ?tk (Ge (CN 0 c e)))"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3014
      using pos_divide_le_eq[OF kpos,
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3015
          where a="real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e" and b="0", symmetric]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3016
        real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3017
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3018
      by (simp add: ti)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3019
    finally have ?case . }
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3020
  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
  3021
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3022
  case (9 i c e)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3023
  then have cp: "c > 0" and nb: "numbound0 e" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3024
  { assume kdc: "k dvd c"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3025
    from tint have ti: "real_of_int \<lfloor>?N (real_of_int x) t\<rfloor> = ?N (real_of_int x) t" using isint_def by simp
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  3026
    from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3027
      numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3028
      numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) }
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3029
  moreover
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3030
  { assume *: "\<not> k dvd c"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3031
    from kpos have knz: "k\<noteq>0" by simp hence knz': "real_of_int k \<noteq> 0" by simp
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3032
    from tint have ti: "real_of_int \<lfloor>?N (real_of_int x) t\<rfloor> = ?N (real_of_int x) t" using isint_def by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3033
    from assms * have "?I (real_of_int x) (?s (Dvd i (CN 0 c e))) = (real_of_int i * real_of_int k rdvd (real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k)"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  3034
      using real_of_int_div[OF kdt]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3035
        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3036
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3037
      by (simp add: ti algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3038
    also have "\<dots> = (?I ?tk (Dvd i (CN 0 c e)))"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3039
      using rdvd_mult[OF knz, where n="i"]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3040
        real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3041
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3042
      by (simp add: ti)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3043
    finally have ?case . }
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3044
  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
  3045
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3046
  case (10 i c e)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3047
  then have cp: "c > 0" and nb: "numbound0 e" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3048
  { assume kdc: "k dvd c"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3049
    from tint have ti: "real_of_int \<lfloor>?N (real_of_int x) t\<rfloor> = ?N (real_of_int x) t" using isint_def by simp
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  3050
    from kdc have  ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3051
      numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3052
      numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) }
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3053
  moreover
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3054
  { assume *: "\<not> k dvd c"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3055
    from kpos have knz: "k\<noteq>0" by simp hence knz': "real_of_int k \<noteq> 0" by simp
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3056
    from tint have ti: "real_of_int \<lfloor>?N (real_of_int x) t\<rfloor> = ?N (real_of_int x) t" using isint_def by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3057
    from assms * have "?I (real_of_int x) (?s (NDvd i (CN 0 c e))) = (\<not> (real_of_int i * real_of_int k rdvd (real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k))"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  3058
      using real_of_int_div[OF kdt]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3059
        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3060
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3061
      by (simp add: ti algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3062
    also have "\<dots> = (?I ?tk (NDvd i (CN 0 c e)))"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  3063
      using rdvd_mult[OF knz, where n="i"] real_of_int_div[OF kdt]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3064
        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3065
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3066
      by (simp add: ti)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3067
    finally have ?case . }
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3068
  ultimately show ?case by blast
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3069
qed (simp_all add: bound0_I[where bs="bs" and b="real_of_int (\<lfloor>?N b' t\<rfloor> div k)" and b'="real_of_int x"]
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3070
  numbound0_I[where bs="bs" and b="real_of_int (\<lfloor>?N b' t\<rfloor> div k)" and b'="real_of_int x"])
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  3071
23264
324622260d29 Added 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
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3073
lemma \<sigma>_\<rho>_nb: assumes lp:"iszlfm p (a#bs)" and nb: "numbound0 t"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3074
  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
  3075
  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
  3076
  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
  3077
324622260d29 Added 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
lemma \<rho>_l:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3079
  assumes lp: "iszlfm p (real_of_int (i::int)#bs)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3080
  shows "\<forall> (b,k) \<in> set (\<rho> p). k >0 \<and> numbound0 b \<and> isint b (real_of_int 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
  3081
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
  3082
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3083
lemma \<alpha>_\<rho>_l:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3084
  assumes lp: "iszlfm p (real_of_int (i::int)#bs)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3085
  shows "\<forall> (b,k) \<in> set (\<alpha>_\<rho> p). k >0 \<and> numbound0 b \<and> isint b (real_of_int i#bs)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3086
using lp isint_add [OF isint_c[where j="- 1"],where bs="real_of_int i#bs"]
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3087
 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
  3088
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3089
lemma \<rho>: assumes lp: "iszlfm p (real_of_int (i::int) #bs)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3090
  and pi: "Ifm (real_of_int i#bs) p"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3091
  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
  3092
  and dp: "d > 0"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3093
  and nob: "\<forall>(e,c) \<in> set (\<rho> p). \<forall> j\<in> {1 .. c*d}. real_of_int (c*i) \<noteq> Inum (real_of_int i#bs) e + real_of_int 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
  3094
  (is "\<forall>(e,c) \<in> set (\<rho> p). \<forall> j\<in> {1 .. c*d}. _ \<noteq> ?N i e + _")
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3095
  shows "Ifm (real_of_int(i - d)#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
  3096
  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
  3097
proof(induct p rule: iszlfm.induct)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3098
  case (3 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real_of_int i#bs)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3099
    and pi: "real_of_int (c*i) = - 1 -  ?N i e + real_of_int (1::int)" and nob: "\<forall> j\<in> {1 .. c*d}. real_of_int (c*i) \<noteq> -1 - ?N i e + real_of_int 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
  3100
    by simp+
324622260d29 Added 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
  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
  3102
  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
  3103
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3104
  case (4 c e)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3105
  hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real_of_int i#bs)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3106
    and nob: "\<forall> j\<in> {1 .. c*d}. real_of_int (c*i) \<noteq> - ?N i e + real_of_int 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
  3107
    by simp+
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3108
  {assume "real_of_int (c*i) \<noteq> - ?N i e + real_of_int (c*d)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3109
    with numbound0_I[OF nb, where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"]
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  3110
    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
  3111
  moreover
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3112
  {assume pi: "real_of_int (c*i) = - ?N i e + real_of_int (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
  3113
    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
  3114
    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
  3115
  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
  3116
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
  3117
  case (5 c e) hence cp: "c > 0" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3118
  from 5 mult_strict_left_mono[OF dp cp, simplified of_int_less_iff[symmetric]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3119
    of_int_mult]
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3120
  show ?case using 5 dp
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3121
    apply (simp add: numbound0_I[where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"]
56544
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56479
diff changeset
  3122
      algebra_simps del: mult_pos_pos)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3123
     by (metis add.right_neutral of_int_0_less_iff of_int_mult pos_add_strict)
23264
324622260d29 Added 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
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3125
  case (6 c e) hence cp: "c > 0" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3126
  from 6 mult_strict_left_mono[OF dp cp, simplified of_int_less_iff[symmetric]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3127
    of_int_mult]
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3128
  show ?case using 6 dp
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3129
    apply (simp add: numbound0_I[where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"]
56544
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56479
diff changeset
  3130
      algebra_simps del: mult_pos_pos)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3131
      using order_trans by fastforce
23264
324622260d29 Added 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
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3133
  case (7 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real_of_int i#bs)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3134
    and nob: "\<forall> j\<in> {1 .. c*d}. real_of_int (c*i) \<noteq> - ?N i e + real_of_int j"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3135
    and pi: "real_of_int (c*i) + ?N i e > 0" and cp': "real_of_int 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
  3136
    by simp+
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3137
  let ?fe = "\<lfloor>?N i e\<rfloor>"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3138
  from pi cp have th:"(real_of_int i +?N i e / real_of_int c)*real_of_int c > 0" by (simp add: algebra_simps)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3139
  from pi ei[simplified isint_iff] have "real_of_int (c*i + ?fe) > real_of_int (0::int)" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3140
  hence pi': "c*i + ?fe > 0" by (simp only: of_int_less_iff[symmetric])
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3141
  have "real_of_int (c*i) + ?N i e > real_of_int (c*d) \<or> real_of_int (c*i) + ?N i e \<le> real_of_int (c*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
  3142
  moreover
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3143
  {assume "real_of_int (c*i) + ?N i e > real_of_int (c*d)" hence ?case
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3144
      by (simp add: algebra_simps
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3145
        numbound0_I[OF nb,where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"])}
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3146
  moreover
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3147
  {assume H:"real_of_int (c*i) + ?N i e \<le> real_of_int (c*d)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3148
    with ei[simplified isint_iff] have "real_of_int (c*i + ?fe) \<le> real_of_int (c*d)" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3149
    hence pid: "c*i + ?fe \<le> c*d" by (simp only: of_int_le_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
  3150
    with pi' have "\<exists> j1\<in> {1 .. c*d}. c*i + ?fe = j1" by auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3151
    hence "\<exists> j1\<in> {1 .. c*d}. real_of_int (c*i) = - ?N i e + real_of_int j1"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3152
      unfolding Bex_def using ei[simplified isint_iff] by fastforce
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3153
    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
  3154
  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
  3155
next
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3156
  case (8 c e)  hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real_of_int i#bs)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3157
    and nob: "\<forall> j\<in> {1 .. c*d}. real_of_int (c*i) \<noteq> - 1 - ?N i e + real_of_int j"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3158
    and pi: "real_of_int (c*i) + ?N i e \<ge> 0" and cp': "real_of_int 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
  3159
    by simp+
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3160
  let ?fe = "\<lfloor>?N i e\<rfloor>"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3161
  from pi cp have th:"(real_of_int i +?N i e / real_of_int c)*real_of_int c \<ge> 0" by (simp add: algebra_simps)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3162
  from pi ei[simplified isint_iff] have "real_of_int (c*i + ?fe) \<ge> real_of_int (0::int)" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3163
  hence pi': "c*i + 1 + ?fe \<ge> 1" by (simp only: of_int_le_iff[symmetric])
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3164
  have "real_of_int (c*i) + ?N i e \<ge> real_of_int (c*d) \<or> real_of_int (c*i) + ?N i e < real_of_int (c*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
  3165
  moreover
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3166
  {assume "real_of_int (c*i) + ?N i e \<ge> real_of_int (c*d)" hence ?case
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3167
      by (simp add: algebra_simps
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3168
        numbound0_I[OF nb,where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"])}
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3169
  moreover
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3170
  {assume H:"real_of_int (c*i) + ?N i e < real_of_int (c*d)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3171
    with ei[simplified isint_iff] have "real_of_int (c*i + ?fe) < real_of_int (c*d)" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3172
    hence pid: "c*i + 1 + ?fe \<le> c*d" by (simp only: of_int_le_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
  3173
    with pi' have "\<exists> j1\<in> {1 .. c*d}. c*i + 1+ ?fe = j1" by auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3174
    hence "\<exists> j1\<in> {1 .. c*d}. real_of_int (c*i) + 1= - ?N i e + real_of_int j1"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3175
      unfolding Bex_def using ei[simplified isint_iff] by fastforce
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3176
    hence "\<exists> j1\<in> {1 .. c*d}. real_of_int (c*i) = (- ?N i e + real_of_int j1) - 1"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3177
      by (simp only: algebra_simps)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3178
        hence "\<exists> j1\<in> {1 .. c*d}. real_of_int (c*i) = - 1 - ?N i e + real_of_int j1"
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54230
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
    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
  3181
  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
  3182
next
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3183
  case (9 j c e)  hence p: "real_of_int j rdvd real_of_int (c*i) + ?N i e" (is "?p x") and cp: "c > 0" and bn:"numbound0 e"  by simp+
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3184
  let ?e = "Inum (real_of_int i # bs) e"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3185
  from 9 have "isint e (real_of_int i #bs)"  by simp
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3186
  hence ie: "real_of_int \<lfloor>?e\<rfloor> = ?e" using isint_iff[where n="e" and bs="(real_of_int i)#bs"] numbound0_I[OF bn,where b="real_of_int i" and b'="real_of_int i" and bs="bs"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3187
    by (simp add: isint_iff)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3188
  from 9 have id: "j dvd d" by simp
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3189
  from ie[symmetric] have "?p i = (real_of_int j rdvd real_of_int (c*i + \<lfloor>?e\<rfloor>))" by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3190
  also have "\<dots> = (j dvd c*i + \<lfloor>?e\<rfloor>)"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3191
    using int_rdvd_iff [where i="j" and t="c*i + \<lfloor>?e\<rfloor>"] by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3192
  also have "\<dots> = (j dvd c*i - c*d + \<lfloor>?e\<rfloor>)"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3193
    using dvd_period[OF id, where x="c*i" and c="-c" and t="\<lfloor>?e\<rfloor>"] by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3194
  also have "\<dots> = (real_of_int j rdvd real_of_int (c*i - c*d + \<lfloor>?e\<rfloor>))"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3195
    using int_rdvd_iff[where i="j" and t="(c*i - c*d + \<lfloor>?e\<rfloor>)",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
  3196
      ie by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3197
  also have "\<dots> = (real_of_int j rdvd real_of_int (c*(i - d)) + ?e)"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3198
    using ie by (simp add:algebra_simps)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3199
  finally show ?case
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3200
    using numbound0_I[OF bn,where b="real_of_int i - real_of_int d" and b'="real_of_int i" and bs="bs"] p
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3201
    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
  3202
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3203
  case (10 j c e)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3204
  hence p: "\<not> (real_of_int j rdvd real_of_int (c*i) + ?N i e)" (is "?p x") and cp: "c > 0" and bn:"numbound0 e"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3205
    by simp+
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3206
  let ?e = "Inum (real_of_int i # bs) e"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3207
  from 10 have "isint e (real_of_int i #bs)"  by simp
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3208
  hence ie: "real_of_int \<lfloor>?e\<rfloor> = ?e"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3209
    using isint_iff[where n="e" and bs="(real_of_int i)#bs"] numbound0_I[OF bn,where b="real_of_int i" and b'="real_of_int i" and bs="bs"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3210
    by (simp add: isint_iff)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3211
  from 10 have id: "j dvd d" by simp
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3212
  from ie[symmetric] have "?p i = (\<not> (real_of_int j rdvd real_of_int (c*i + \<lfloor>?e\<rfloor>)))" by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3213
  also have "\<dots> = Not (j dvd c*i + \<lfloor>?e\<rfloor>)"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3214
    using int_rdvd_iff [where i="j" and t="c*i + \<lfloor>?e\<rfloor>"] by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3215
  also have "\<dots> = Not (j dvd c*i - c*d + \<lfloor>?e\<rfloor>)"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3216
    using dvd_period[OF id, where x="c*i" and c="-c" and t="\<lfloor>?e\<rfloor>"] by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3217
  also have "\<dots> = Not (real_of_int j rdvd real_of_int (c*i - c*d + \<lfloor>?e\<rfloor>))"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3218
    using int_rdvd_iff[where i="j" and t="(c*i - c*d + \<lfloor>?e\<rfloor>)",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
  3219
      ie by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3220
  also have "\<dots> = Not (real_of_int j rdvd real_of_int (c*(i - d)) + ?e)"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3221
    using ie by (simp add:algebra_simps)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3222
  finally show ?case
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3223
    using numbound0_I[OF bn,where b="real_of_int i - real_of_int d" and b'="real_of_int i" and bs="bs"] p
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3224
    by (simp add: algebra_simps)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3225
qed (auto simp add: numbound0_I[where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int 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
  3226
324622260d29 Added 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
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
  3228
  shows "bound0 (\<sigma> p k t)"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3229
  using \<sigma>_\<rho>_nb[OF lp nb] nb by (simp add: \<sigma>_def)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3230
23264
324622260d29 Added 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
lemma \<rho>':   assumes lp: "iszlfm p (a #bs)"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3232
  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
  3233
  and dp: "d > 0"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3234
  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_of_int x#bs) p \<longrightarrow> Ifm (real_of_int (x - d)#bs) p" (is "\<forall> x. ?b x \<longrightarrow> ?P x \<longrightarrow> ?P (x - 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
  3235
proof(clarify)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3236
  fix x
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3237
  assume nob1:"?b x" and px: "?P x"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3238
  from iszlfm_gen[OF lp, rule_format, where y="real_of_int x"] have lp': "iszlfm p (real_of_int x#bs)".
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3239
  have nob: "\<forall>(e, c)\<in>set (\<rho> p). \<forall>j\<in>{1..c * d}. real_of_int (c * x) \<noteq> Inum (real_of_int x # bs) e + real_of_int 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
  3240
  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
  3241
    fix e c j assume ecR: "(e,c) \<in> set (\<rho> p)" and jD: "j\<in> {1 .. c*d}"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3242
      and cx: "real_of_int (c*x) = Inum (real_of_int x#bs) e + real_of_int j"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3243
    let ?e = "Inum (real_of_int x#bs) e"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3244
    from \<rho>_l[OF lp'] ecR have ei:"isint e (real_of_int x#bs)" and cp:"c>0" and nb:"numbound0 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
  3245
      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
  3246
    from numbound0_gen [OF nb ei, rule_format,where y="a"] have "isint e (a#bs)" .
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3247
    from cx ei[simplified isint_iff] have "real_of_int (c*x) = real_of_int (\<lfloor>?e\<rfloor> + j)" by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3248
    hence cx: "c*x = \<lfloor>?e\<rfloor> + j" by (simp only: of_int_eq_iff)
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3249
    hence cdej:"c dvd \<lfloor>?e\<rfloor> + j" by (simp add: dvd_def) (rule_tac x="x" in exI, simp)
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3250
    hence "real_of_int c rdvd real_of_int (\<lfloor>?e\<rfloor> + j)" by (simp only: int_rdvd_iff)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3251
    hence rcdej: "real_of_int c rdvd ?e + real_of_int j" by (simp add: ei[simplified isint_iff])
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3252
    from cx have "(c*x) div c = (\<lfloor>?e\<rfloor> + j) div c" by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3253
    with cp have "x = (\<lfloor>?e\<rfloor> + j) div c" by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3254
    with px have th: "?P ((\<lfloor>?e\<rfloor> + j) div c)" by auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3255
    from cp have cp': "real_of_int c > 0" by simp
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3256
    from cdej have cdej': "c dvd \<lfloor>Inum (real_of_int x#bs) (Add e (C j))\<rfloor>" 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
  3257
    from nb have nb': "numbound0 (Add e (C j))" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3258
    have ji: "isint (C j) (real_of_int x#bs)" by (simp add: isint_def)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3259
    from isint_add[OF ei ji] have ei':"isint (Add e (C j)) (real_of_int x#bs)" .
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3260
    from th \<sigma>_\<rho>[where b'="real_of_int x", OF lp' cp' nb' ei' cdej',symmetric]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3261
    have "Ifm (real_of_int x#bs) (\<sigma>_\<rho> p (Add e (C j), c))" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3262
    with rcdej have th: "Ifm (real_of_int x#bs) (\<sigma> p c (Add e (C j)))" by (simp add: \<sigma>_def)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3263
    from th bound0_I[OF \<sigma>_nb[OF lp nb', where k="c"],where bs="bs" and b="real_of_int x" and b'="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
  3264
    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
  3265
      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
  3266
  qed
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3267
  from \<rho>[OF lp' px d dp nob] show "?P (x -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
  3268
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
  3269
324622260d29 Added 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
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3271
lemma rl_thm:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3272
  assumes lp: "iszlfm p (real_of_int (i::int)#bs)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3273
  shows "(\<exists> (x::int). Ifm (real_of_int x#bs) p) = ((\<exists> j\<in> {1 .. \<delta> p}. Ifm (real_of_int 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)))))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3274
  (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))"
23264
324622260d29 Added 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
    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
  3276
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
  3277
  let ?d= "\<delta> p"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3278
  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
  3279
  { 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
  3280
    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
  3281
  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
  3282
  { fix e c j assume exR:"(e,c) \<in> ?R" and jD:"j\<in> {1 .. c*?d}" and spx:"?SP c e j"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3283
    from exR \<rho>_l[OF lp] have nb: "numbound0 e" and ei:"isint e (real_of_int i#bs)" and cp: "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
  3284
      by auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3285
    have "isint (C j) (real_of_int i#bs)" by (simp add: isint_iff)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3286
    with isint_add[OF numbound0_gen[OF nb ei,rule_format, where y="real_of_int i"]]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3287
    have eji:"isint (Add e (C j)) (real_of_int i#bs)" 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
  3288
    from nb have nb': "numbound0 (Add e (C j))" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3289
    from spx bound0_I[OF \<sigma>_nb[OF lp nb', where k="c"], where bs="bs" and b="a" and b'="real_of_int i"]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3290
    have spx': "Ifm (real_of_int i # bs) (\<sigma> p c (Add e (C j)))" by blast
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3291
    from spx' have rcdej:"real_of_int c rdvd (Inum (real_of_int i#bs) (Add e (C j)))"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3292
      and sr:"Ifm (real_of_int i#bs) (\<sigma>_\<rho> p (Add e (C j),c))" by (simp add: \<sigma>_def)+
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3293
    from rcdej eji[simplified isint_iff]
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3294
    have "real_of_int c rdvd real_of_int \<lfloor>Inum (real_of_int i#bs) (Add e (C j))\<rfloor>" by simp
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3295
    hence cdej:"c dvd \<lfloor>Inum (real_of_int i#bs) (Add e (C j))\<rfloor>" by (simp only: int_rdvd_iff)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3296
    from cp have cp': "real_of_int c > 0" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3297
    from \<sigma>_\<rho>[OF lp cp' nb' eji cdej] spx' have "?P (\<lfloor>Inum (real_of_int 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
  3298
      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
  3299
    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
  3300
    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
  3301
  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
  3302
  { 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
  3303
    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
  3304
    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
  3305
    from minusinf_inf[OF lp] obtain z where z:"\<forall> x<z. ?MP x = ?P x" by blast
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3306
    have zp: "\<bar>x - z\<bar> + 1 \<ge> 0" by arith
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3307
    from decr_lemma[OF dp,where x="x" and z="z"]
23264
324622260d29 Added 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
      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
  3309
    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
  3310
  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
  3311
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
  3312
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3313
lemma mirror_\<alpha>_\<rho>:   assumes lp: "iszlfm p (a#bs)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3314
  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
  3315
  using lp
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3316
  by (induct p rule: mirror.induct) (simp_all add: split_def image_Un)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3317
61586
5197a2ecb658 isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
  3318
text \<open>The \<open>\<real>\<close> part\<close>
5197a2ecb658 isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
  3319
5197a2ecb658 isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
  3320
text\<open>Linearity for fm where Bound 0 ranges over \<open>\<real>\<close>\<close>
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3321
fun isrlfm :: "fm \<Rightarrow> bool"   (* Linearity test for fm *)
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3322
where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3323
  "isrlfm (And p q) = (isrlfm p \<and> isrlfm q)"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3324
| "isrlfm (Or p q) = (isrlfm p \<and> isrlfm q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3325
| "isrlfm (Eq  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3326
| "isrlfm (NEq (CN 0 c e)) = (c>0 \<and> numbound0 e)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3327
| "isrlfm (Lt  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3328
| "isrlfm (Le  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3329
| "isrlfm (Gt  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3330
| "isrlfm (Ge  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3331
| "isrlfm p = (isatom p \<and> (bound0 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
  3332
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  3333
definition fp :: "fm \<Rightarrow> int \<Rightarrow> num \<Rightarrow> int \<Rightarrow> fm" where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3334
  "fp p n s j \<equiv> (if n > 0 then
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3335
            (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
  3336
                        (Lt (CN 0 n (Sub s (Add (Floor s) (C (j+1))))))))
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3337
            else
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3338
            (And p (And (Le (CN 0 (-n) (Add (Neg s) (Add (Floor s) (C 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
  3339
                        (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
  3340
324622260d29 Added 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
  (* splits the bounded from the unbounded part*)
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3342
fun rsplit0 :: "num \<Rightarrow> (fm \<times> int \<times> num) list"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3343
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
  3344
  "rsplit0 (Bound 0) = [(T,1,C 0)]"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3345
| "rsplit0 (Add a b) = (let acs = rsplit0 a ; bcs = rsplit0 b
24336
fff40259f336 removed allpairs
nipkow
parents: 24249
diff changeset
  3346
              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
  3347
| "rsplit0 (Sub a b) = rsplit0 (Add a (Neg b))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  3348
| "rsplit0 (Neg a) = map (\<lambda> (p,n,s). (p,-n,Neg s)) (rsplit0 a)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3349
| "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
  3350
      (\<lambda> (p,n,s). if n=0 then [(p,0,Floor s)]
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3351
          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
  3352
       (rsplit0 a))"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  3353
| "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
  3354
| "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
  3355
| "rsplit0 (CF c t s) = rsplit0 (Add (Mul c (Floor t)) s)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  3356
| "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
  3357
| "rsplit0 t = [(T,0,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
  3358
324622260d29 Added 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
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
  3360
  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
  3361
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
  3362
  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
  3363
324622260d29 Added 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
324622260d29 Added 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
lemma rsplit0_cs:
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3366
  shows "\<forall> (p,n,s) \<in> set (rsplit0 t).
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3367
  (Ifm (x#bs) p \<longrightarrow>  (Inum (x#bs) t = Inum (x#bs) (CN 0 n s))) \<and> numbound0 s \<and> isrlfm 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
  3368
  (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
  3369
proof(induct t rule: rsplit0.induct)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3370
  case (5 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
  3371
  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
  3372
  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
  3373
  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
  3374
  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
  3375
  have int_cases: "\<forall> (i::int). i= 0 \<or> i < 0 \<or> i > 0" by arith
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3376
  have U1: "(UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set (?ff (p,n,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
  3377
    (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set [(p,0,Floor s)]))" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3378
  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
  3379
    ?ff (p,n,s) = map (?f(p,n,s)) [0..n]" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3380
  hence U2: "(UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) =
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3381
    (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
  3382
    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
  3383
  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
  3384
    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
  3385
    assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c"
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69266
diff changeset
  3386
    thus "(\<Union>(a, b, c)\<in>M. set (f (a, b, c))) = (\<Union>(a, b, c)\<in>M. set (g a b 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
  3387
      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
  3388
  qed
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3389
  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
  3390
    by auto
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69266
diff changeset
  3391
  hence U3: "(\<Union> ((\<lambda>(p,n,s). set (?ff (p,n,s))) ` {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0})) =
b021008c5397 removed legacy input syntax
haftmann
parents: 69266
diff changeset
  3392
    (\<Union> ((\<lambda>(p,n,s). set (map (?f(p,n,s)) [n..0])) ` {(p,n,s). (p,n,s)\<in> ?SS a\<and>n<0}))"
b021008c5397 removed legacy input syntax
haftmann
parents: 69266
diff changeset
  3393
  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
  3394
    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
  3395
    assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c"
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69266
diff changeset
  3396
    thus "(\<Union>(a, b, c)\<in>M. set (f (a, b, c))) = (\<Union>(a, b, c)\<in>M. set (g a b 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
  3397
      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
  3398
  qed
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69266
diff changeset
  3399
  have "?SS (Floor a) = \<Union> ((\<lambda>x. set (?ff x)) ` ?SS a)"
46130
4821af078cd6 prefer concat over foldl append []
haftmann
parents: 45740
diff changeset
  3400
    by auto
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69266
diff changeset
  3401
  also have "\<dots> = \<Union> ((\<lambda> (p,n,s). set (?ff (p,n,s))) ` ?SS a)"
b021008c5397 removed legacy input syntax
haftmann
parents: 69266
diff changeset
  3402
    by blast
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3403
  also have "\<dots> =
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3404
    ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3405
    (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3406
    (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (?ff (p,n,s)))))"
68270
2bc921b2159b treat gcd_eq_1_imp_coprime analogously to mod_0_imp_dvd
haftmann
parents: 67613
diff changeset
  3407
    by (auto split: if_splits)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3408
  also have "\<dots> =
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3409
    ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set [(p,0,Floor s)])) Un
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3410
   (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
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3411
   (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
  3412
    set (map (?f(p,n,s)) [n..0]))))" by (simp only: U1 U2 U3)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3413
  also have "\<dots> =
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3414
    ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3415
    (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). (?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
  3416
    (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {n .. 0})))"
57816
d8bbb97689d3 no need for 'set_simps' now that 'datatype_new' generates the desired 'set' property
blanchet
parents: 57514
diff changeset
  3417
    by (simp only: set_map set_upto list.set)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3418
  also have "\<dots> =
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3419
    ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3420
    (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
68270
2bc921b2159b treat gcd_eq_1_imp_coprime analogously to mod_0_imp_dvd
haftmann
parents: 67613
diff changeset
  3421
    (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}})))"
2bc921b2159b treat gcd_eq_1_imp_coprime analogously to mod_0_imp_dvd
haftmann
parents: 67613
diff changeset
  3422
    by blast
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3423
  finally
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3424
  have FS: "?SS (Floor a) =
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3425
    ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3426
    (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
68270
2bc921b2159b treat gcd_eq_1_imp_coprime analogously to mod_0_imp_dvd
haftmann
parents: 67613
diff changeset
  3427
    (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}})))"
2bc921b2159b treat gcd_eq_1_imp_coprime analogously to mod_0_imp_dvd
haftmann
parents: 67613
diff changeset
  3428
    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
  3429
  show ?case
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3430
  proof(simp only: FS, clarsimp simp del: Ifm.simps Inum.simps, -)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3431
    fix p n s
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3432
    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
  3433
    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
  3434
       (\<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
  3435
           (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
  3436
           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
  3437
           (\<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
  3438
                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
  3439
       (\<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
  3440
           (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
  3441
           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
  3442
           (\<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
  3443
                n = 0 \<and> s = Add (Floor ba) (C j) \<and> ac \<le> j \<and> j \<le> 0))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3444
    moreover
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3445
    { fix s'
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3446
      assume "(p, 0, s') \<in> ?SS a" and "n = 0" and "s = Floor s'"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3447
      hence ?ths using 5(1) by auto }
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3448
    moreover
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3449
    { fix p' n' s' j
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3450
      assume pns: "(p', n', s') \<in> ?SS a"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3451
        and np: "0 < n'"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3452
        and p_def: "p = ?p (p',n',s') j"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3453
        and n0: "n = 0"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3454
        and s_def: "s = (Add (Floor s') (C j))"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3455
        and jp: "0 \<le> j" and jn: "j \<le> n'"
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60533
diff changeset
  3456
      from 5 pns have H:"(Ifm ((x::real) # (bs::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
  3457
          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
  3458
          numbound0 s' \<and> isrlfm p'" by blast
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3459
      hence nb: "numbound0 s'" by simp
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3460
      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
  3461
      let ?nxs = "CN 0 n' s'"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3462
      let ?l = "\<lfloor>?N s'\<rfloor> + j"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3463
      from H
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3464
      have "?I (?p (p',n',s') j) \<longrightarrow>
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3465
          (((?N ?nxs \<ge> real_of_int ?l) \<and> (?N ?nxs < real_of_int (?l + 1))) \<and> (?N a = ?N ?nxs ))"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3466
        by (simp add: fp_def np algebra_simps)
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3467
      also have "\<dots> \<longrightarrow> \<lfloor>?N ?nxs\<rfloor> = ?l \<and> ?N a = ?N ?nxs"
66515
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66453
diff changeset
  3468
        using floor_eq_iff[where x="?N ?nxs" and a="?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
  3469
      moreover
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3470
      have "\<dots> \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3471
      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
  3472
        by blast
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3473
      with s_def n0 p_def nb nf have ?ths by auto}
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3474
    moreover
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3475
    { fix p' n' s' j
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3476
      assume pns: "(p', n', s') \<in> ?SS a"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3477
        and np: "n' < 0"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3478
        and p_def: "p = ?p (p',n',s') j"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3479
        and n0: "n = 0"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3480
        and s_def: "s = (Add (Floor s') (C j))"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3481
        and jp: "n' \<le> j" and jn: "j \<le> 0"
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60533
diff changeset
  3482
      from 5 pns have H:"(Ifm ((x::real) # (bs::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
  3483
          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
  3484
          numbound0 s' \<and> isrlfm p'" by blast
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3485
      hence nb: "numbound0 s'" by simp
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3486
      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
  3487
      let ?nxs = "CN 0 n' s'"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3488
      let ?l = "\<lfloor>?N s'\<rfloor> + j"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3489
      from H
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3490
      have "?I (?p (p',n',s') j) \<longrightarrow>
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3491
          (((?N ?nxs \<ge> real_of_int ?l) \<and> (?N ?nxs < real_of_int (?l + 1))) \<and> (?N a = ?N ?nxs ))"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3492
        by (simp add: np fp_def algebra_simps)
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3493
      also have "\<dots> \<longrightarrow> \<lfloor>?N ?nxs\<rfloor> = ?l \<and> ?N a = ?N ?nxs"
66515
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66453
diff changeset
  3494
        using floor_eq_iff[where x="?N ?nxs" and a="?l"] by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3495
      moreover
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3496
      have "\<dots> \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))"  by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3497
      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
  3498
        by blast
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3499
      with s_def n0 p_def nb nf have ?ths by auto}
61652
90c65a811257 MIR decision procedure again working
paulson <lp15@cam.ac.uk>
parents: 61649
diff changeset
  3500
    ultimately show ?ths by fastforce
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3501
  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
  3502
next
28741
1b257449f804 simproc for let
haftmann
parents: 28290
diff changeset
  3503
  case (3 a b) then show ?case
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53168
diff changeset
  3504
    by auto
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3505
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
  3506
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3507
lemma real_in_int_intervals:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3508
  assumes xb: "real_of_int m \<le> x \<and> x < real_of_int ((n::int) + 1)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3509
  shows "\<exists> j\<in> {m.. n}. real_of_int j \<le> x \<and> x < real_of_int (j+1)" (is "\<exists> j\<in> ?N. ?P j")
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3510
by (rule bexI[where P="?P" and x="\<lfloor>x\<rfloor>" and A="?N"])
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3511
(auto simp add: floor_less_iff[where x="x" and z="n+1", simplified]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3512
  xb[simplified] floor_mono[where x="real_of_int m" and y="x", OF conjunct1[OF xb], simplified floor_of_int[where z="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
  3513
324622260d29 Added 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
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
  3515
  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
  3516
  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
  3517
proof(induct t rule: rsplit0.induct)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3518
  case (2 a b)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3519
  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
  3520
  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
  3521
  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
  3522
  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
  3523
  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
  3524
    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
  3525
  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
  3526
  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
  3527
    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
  3528
  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
  3529
  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
  3530
  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
  3531
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3532
  case (5 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
  3533
  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
  3534
  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
  3535
  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
  3536
  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
  3537
  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
  3538
  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
  3539
  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
  3540
    by auto
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3541
  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
  3542
  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
  3543
    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
  3544
    assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c"
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69266
diff changeset
  3545
    thus "(\<Union>(a, b, c)\<in>M. set (f (a, b, c))) = (\<Union>(a, b, c)\<in>M. set (g a b 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
  3546
      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
  3547
  qed
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3548
  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
  3549
    by auto
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3550
  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
  3551
  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
  3552
    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
  3553
    assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c"
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69266
diff changeset
  3554
    thus "(\<Union>(a, b, c)\<in>M. set (f (a, b, c))) = (\<Union>(a, b, c)\<in>M. set (g a b 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
  3555
      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
  3556
  qed
24473
acd19ea21fbb fixed Proofs
chaieb
parents: 24348
diff changeset
  3557
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69266
diff changeset
  3558
  have "?SS (Floor a) = \<Union> ((\<lambda>x. set (?ff x)) ` ?SS a)" by auto
b021008c5397 removed legacy input syntax
haftmann
parents: 69266
diff changeset
  3559
  also have "\<dots> = \<Union> ((\<lambda> (p,n,s). set (?ff (p,n,s))) ` ?SS a)" by blast
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3560
  also have "\<dots> =
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3561
    ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3562
    (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3563
    (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (?ff (p,n,s)))))"
68270
2bc921b2159b treat gcd_eq_1_imp_coprime analogously to mod_0_imp_dvd
haftmann
parents: 67613
diff changeset
  3564
    by (auto split: if_splits)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3565
  also have "\<dots> =
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3566
    ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set [(p,0,Floor s)])) Un
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3567
    (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
  3568
    (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
  3569
    by (simp only: U1 U2 U3)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3570
  also have "\<dots> =
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3571
    ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3572
    (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). (?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
  3573
    (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {n .. 0})))"
57816
d8bbb97689d3 no need for 'set_simps' now that 'datatype_new' generates the desired 'set' property
blanchet
parents: 57514
diff changeset
  3574
    by (simp only: set_map set_upto list.set)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3575
  also have "\<dots> =
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3576
    ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3577
    (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
  3578
    (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
  3579
    by blast
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3580
  finally
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3581
  have FS: "?SS (Floor a) =
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3582
    ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3583
    (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
  3584
    (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
  3585
    by blast
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3586
  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
  3587
  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
  3588
  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
  3589
  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
  3590
    by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3591
23264
324622260d29 Added 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
  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
  3593
  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
  3594
  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
  3595
  {
324622260d29 Added 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
    assume np: "n > 0"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3597
    from of_int_floor_le[of "?N s"] have "?N (Floor s) \<le> ?N s" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3598
    also from mult_left_mono[OF xp] np have "?N s \<le> real_of_int n * x + ?N s" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3599
    finally have "?N (Floor s) \<le> real_of_int n * x + ?N 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
  3600
    moreover
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3601
    {from x1 np have "real_of_int n *x + ?N s < real_of_int n + ?N s" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3602
      also from real_of_int_floor_add_one_gt[where r="?N s"]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3603
      have "\<dots> < real_of_int n + ?N (Floor s) + 1" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3604
      finally have "real_of_int n *x + ?N s < ?N (Floor s) + real_of_int (n+1)" by simp}
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3605
    ultimately have "?N (Floor s) \<le> real_of_int n *x + ?N s\<and> real_of_int n *x + ?N s < ?N (Floor s) + real_of_int (n+1)" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3606
    hence th: "0 \<le> real_of_int n *x + ?N s - ?N (Floor s) \<and> real_of_int n *x + ?N s - ?N (Floor s) < real_of_int (n+1)" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3607
    from real_in_int_intervals th have  "\<exists> j\<in> {0 .. n}. real_of_int j \<le> real_of_int n *x + ?N s - ?N (Floor s)\<and> real_of_int n *x + ?N s - ?N (Floor s) < real_of_int (j+1)" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3608
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3609
    hence "\<exists> j\<in> {0 .. n}. 0 \<le> real_of_int n *x + ?N s - ?N (Floor s) - real_of_int j \<and> real_of_int n *x + ?N s - ?N (Floor s) - real_of_int (j+1) < 0"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3610
      by(simp only: myle[of _ "real_of_int n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real_of_int 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
  3611
    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
  3612
      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
  3613
    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
  3614
    hence "\<exists>x \<in> {?p (p,n,s) j |j. 0\<le> j \<and> j \<le> n }. ?I x" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3615
    hence ?case using pns
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3616
      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
  3617
    (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
  3618
      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
  3619
  }
324622260d29 Added 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
  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
  3621
  { assume nn: "n < 0" hence np: "-n >0" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3622
    from of_int_floor_le[of "?N s"] have "?N (Floor s) + 1 > ?N s" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3623
    moreover from mult_left_mono_neg[OF xp] nn have "?N s \<ge> real_of_int n * x + ?N s" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3624
    ultimately have "?N (Floor s) + 1 > real_of_int n * x + ?N s" 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
  3625
    moreover
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3626
    {from x1 nn have "real_of_int n *x + ?N s \<ge> real_of_int n + ?N s" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3627
      moreover from of_int_floor_le[of "?N s"]  have "real_of_int n + ?N s \<ge> real_of_int n + ?N (Floor s)" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3628
      ultimately have "real_of_int n *x + ?N s \<ge> ?N (Floor s) + real_of_int n"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  3629
        by (simp only: algebra_simps)}
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3630
    ultimately have "?N (Floor s) + real_of_int n \<le> real_of_int n *x + ?N s\<and> real_of_int n *x + ?N s < ?N (Floor s) + real_of_int (1::int)" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3631
    hence th: "real_of_int n \<le> real_of_int n *x + ?N s - ?N (Floor s) \<and> real_of_int n *x + ?N s - ?N (Floor s) < real_of_int (1::int)" 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
  3632
    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
  3633
    have th2: "\<forall> (a::real). (0 \<ge> - a) = (a \<ge> 0)" by auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3634
    from real_in_int_intervals th  have  "\<exists> j\<in> {n .. 0}. real_of_int j \<le> real_of_int n *x + ?N s - ?N (Floor s)\<and> real_of_int n *x + ?N s - ?N (Floor s) < real_of_int (j+1)" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3635
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3636
    hence "\<exists> j\<in> {n .. 0}. 0 \<le> real_of_int n *x + ?N s - ?N (Floor s) - real_of_int j \<and> real_of_int n *x + ?N s - ?N (Floor s) - real_of_int (j+1) < 0"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3637
      by(simp only: myle[of _ "real_of_int n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real_of_int n *x + ?N s - ?N (Floor s)"])
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3638
    hence "\<exists> j\<in> {n .. 0}. 0 \<ge> - (real_of_int n *x + ?N s - ?N (Floor s) - real_of_int j) \<and> - (real_of_int n *x + ?N s - ?N (Floor s) - real_of_int (j+1)) > 0" by (simp only: th1[rule_format] th2[rule_format])
23264
324622260d29 Added 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
    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
  3640
      using pns by (simp add: fp_def nn algebra_simps
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3641
        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
  3642
    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
  3643
    hence "\<exists>x \<in> {?p (p,n,s) j |j. n\<le> j \<and> j \<le> 0 }. ?I x" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3644
    hence ?case using pns
23464
bc2563c37b1a tuned proofs -- avoid implicit prems;
wenzelm
parents: 23316
diff changeset
  3645
      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
  3646
    (rule disjI2, rule disjI2,rule exI [where x="p"],
23464
bc2563c37b1a tuned proofs -- avoid implicit prems;
wenzelm
parents: 23316
diff changeset
  3647
      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
  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
  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
  3650
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
  3651
324622260d29 Added 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
    (* 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
  3653
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  3654
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
  3655
  "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
  3656
324622260d29 Added 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
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
  3658
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
  3659
324622260d29 Added 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
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
  3661
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
  3662
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3663
lemma foldr_disj_map_rlfm:
23264
324622260d29 Added 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
  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
  3665
  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
  3666
  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
  3667
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
  3668
324622260d29 Added 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
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
  3670
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
  3671
324622260d29 Added 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
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
  3673
  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
  3674
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
  3675
  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
  3676
  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
  3677
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
  3678
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3679
lemma rsplit:
23264
324622260d29 Added 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
  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
  3681
  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
  3682
  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
  3683
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
  3684
  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
  3685
  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
  3686
  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
  3687
  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
  3688
  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
  3689
  hence \<phi>: "?I x \<phi>" and fns: "?I x (f n s)" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3690
  from rsplit0_cs[where t="a" and bs="bs" and x="x", rule_format, OF fnsS] \<phi>
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3691
  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
  3692
  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
  3693
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
  3694
  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
  3695
  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
  3696
  assume ga: "?I x (g a)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3697
  from rsplit0_complete[OF xp x1, where bs="bs" and 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
  3698
  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
  3699
  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
  3700
  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
  3701
  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
  3702
  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
  3703
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
  3704
23997
a23d0b4b1c1f Updated proofs;
chaieb
parents: 23993
diff changeset
  3705
definition lt :: "int \<Rightarrow> num \<Rightarrow> fm" where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3706
  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
  3707
                        else (Gt (CN 0 (-c) (Neg t))))"
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3708
23997
a23d0b4b1c1f Updated proofs;
chaieb
parents: 23993
diff changeset
  3709
definition  le :: "int \<Rightarrow> num \<Rightarrow> fm" where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3710
  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
  3711
                        else (Ge (CN 0 (-c) (Neg t))))"
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3712
23997
a23d0b4b1c1f Updated proofs;
chaieb
parents: 23993
diff changeset
  3713
definition  gt :: "int \<Rightarrow> num \<Rightarrow> fm" where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3714
  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
  3715
                        else (Lt (CN 0 (-c) (Neg t))))"
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3716
23997
a23d0b4b1c1f Updated proofs;
chaieb
parents: 23993
diff changeset
  3717
definition  ge :: "int \<Rightarrow> num \<Rightarrow> fm" where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3718
  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
  3719
                        else (Le (CN 0 (-c) (Neg t))))"
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3720
23997
a23d0b4b1c1f Updated proofs;
chaieb
parents: 23993
diff changeset
  3721
definition  eq :: "int \<Rightarrow> num \<Rightarrow> fm" where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3722
  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
  3723
                        else (Eq (CN 0 (-c) (Neg t))))"
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3724
23997
a23d0b4b1c1f Updated proofs;
chaieb
parents: 23993
diff changeset
  3725
definition neq :: "int \<Rightarrow> num \<Rightarrow> fm" where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3726
  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
  3727
                        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
  3728
324622260d29 Added 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
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
  3730
  (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
  3731
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
  3732
  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
  3733
  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
  3734
  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
  3735
  (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
  3736
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
  3737
324622260d29 Added 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 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
  3739
  by (rule rsplit_l[where f="lt" and a="a"], auto simp add: lt_def,
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  3740
    case_tac s, simp_all, rename_tac nat a b, case_tac "nat", 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
  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 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
  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 (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
  3747
  (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
  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
324622260d29 Added 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
lemma le_l: "isrlfm (rsplit le a)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3751
  by (rule rsplit_l[where f="le" and a="a"], auto simp add: le_def)
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  3752
(case_tac s, simp_all, rename_tac nat a b, case_tac "nat",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
  3753
324622260d29 Added 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
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
  3755
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
  3756
  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
  3757
  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
  3758
  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
  3759
  (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
  3760
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
  3761
lemma gt_l: "isrlfm (rsplit gt a)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3762
  by (rule rsplit_l[where f="gt" and a="a"], auto simp add: gt_def)
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  3763
(case_tac s, simp_all, rename_tac nat a b, case_tac "nat", 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
  3764
324622260d29 Added 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
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
  3766
proof(clarify)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3767
  fix a n 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
  3768
  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
  3769
  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
  3770
  (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
  3771
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
  3772
lemma ge_l: "isrlfm (rsplit ge a)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3773
  by (rule rsplit_l[where f="ge" and a="a"], auto simp add: ge_def)
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  3774
(case_tac s, simp_all, rename_tac nat a b, case_tac "nat", 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
  3775
324622260d29 Added 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
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
  3777
proof(clarify)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3778
  fix a n 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
  3779
  assume H: "?N a = ?N (CN 0 n s)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  3780
  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
  3781
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
  3782
lemma eq_l: "isrlfm (rsplit eq a)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3783
  by (rule rsplit_l[where f="eq" and a="a"], auto simp add: eq_def)
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  3784
(case_tac s, simp_all, rename_tac nat a b, case_tac"nat", 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
  3785
324622260d29 Added 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
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
  3787
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
  3788
  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
  3789
  assume H: "?N a = ?N (CN 0 n s)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  3790
  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
  3791
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
  3792
324622260d29 Added 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
lemma neq_l: "isrlfm (rsplit neq a)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3794
  by (rule rsplit_l[where f="neq" and a="a"], auto simp add: neq_def)
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  3795
(case_tac s, simp_all, rename_tac nat a b, case_tac"nat", 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
  3796
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3797
lemma small_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
  3798
  assumes u0:"0 \<le> u" and u1: "u < 1"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3799
  shows "(-u \<le> real_of_int (n::int)) = (0 \<le> 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
  3800
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
  3801
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3802
lemma small_lt:
23264
324622260d29 Added 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
  assumes u0:"0 \<le> u" and u1: "u < 1"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3804
  shows "(real_of_int (n::int) < real_of_int (m::int) - u) = (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
  3805
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
  3806
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3807
lemma rdvd01_cs:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3808
  assumes up: "u \<ge> 0" and u1: "u<1" and np: "real_of_int n > 0"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3809
  shows "(real_of_int (i::int) rdvd real_of_int (n::int) * u - s) = (\<exists> j\<in> {0 .. n - 1}. real_of_int n * u = s - real_of_int \<lfloor>s\<rfloor> + real_of_int j \<and> real_of_int i rdvd real_of_int (j - \<lfloor>s\<rfloor>))" (is "?lhs = ?rhs")
23264
324622260d29 Added 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
proof-
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3811
  let ?ss = "s - real_of_int \<lfloor>s\<rfloor>"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3812
  from real_of_int_floor_add_one_gt[where r="s", simplified myless[of "s"]]
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  3813
    of_int_floor_le  have ss0:"?ss \<ge> 0" and ss1:"?ss < 1" by (auto simp: floor_less_cancel)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3814
  from np have n0: "real_of_int n \<ge> 0" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3815
  from mult_left_mono[OF up n0] mult_strict_left_mono[OF u1 np]
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3816
  have nu0:"real_of_int n * u - s \<ge> -s" and nun:"real_of_int n * u -s < real_of_int n - s" by auto
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3817
  from int_rdvd_real[where i="i" and x="real_of_int (n::int) * u - s"]
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3818
  have "real_of_int i rdvd real_of_int n * u - s =
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3819
    (i dvd \<lfloor>real_of_int n * u - s\<rfloor> \<and> (real_of_int \<lfloor>real_of_int n * u - s\<rfloor> = real_of_int n * u - 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
  3820
    (is "_ = (?DE)" is "_ = (?D \<and> ?E)") by simp
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3821
  also have "\<dots> = (?DE \<and> real_of_int (\<lfloor>real_of_int n * u - s\<rfloor> + \<lfloor>s\<rfloor>) \<ge> -?ss
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3822
    \<and> real_of_int (\<lfloor>real_of_int n * u - s\<rfloor> + \<lfloor>s\<rfloor>) < real_of_int n - ?ss)" (is "_=(?DE \<and>real_of_int ?a \<ge> _ \<and> real_of_int ?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
  3823
    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
  3824
  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
  3825
  also have "\<dots> = (?DE \<and> (\<exists> j\<in> {0 .. (n - 1)}. ?a = j))" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3826
  also have "\<dots> = (?DE \<and> (\<exists> j\<in> {0 .. (n - 1)}. real_of_int (\<lfloor>real_of_int n * u - s\<rfloor>) = real_of_int j - real_of_int \<lfloor>s\<rfloor> ))"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3827
    by (simp only: algebra_simps of_int_diff[symmetric] of_int_eq_iff)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3828
  also have "\<dots> = ((\<exists> j\<in> {0 .. (n - 1)}. real_of_int n * u - s = real_of_int j - real_of_int \<lfloor>s\<rfloor> \<and> real_of_int i rdvd real_of_int n * u - s))" using int_rdvd_iff[where i="i" and t="\<lfloor>real_of_int n * u - s\<rfloor>"]
23264
324622260d29 Added 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
    by (auto cong: conj_cong)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  3830
  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
  3831
  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
  3832
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
  3833
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3834
definition
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3835
  DVDJ:: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm"
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3836
where
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3837
  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
  3838
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3839
definition
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3840
  NDVDJ:: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm"
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3841
where
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3842
  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
  3843
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3844
lemma DVDJ_DVD:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3845
  assumes xp:"x\<ge> 0" and x1: "x < 1" and np:"real_of_int 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
  3846
  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
  3847
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
  3848
  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
  3849
  let ?s= "Inum (x#bs) s"
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3850
  from foldr_disj_map[where xs="[0..n - 1]" and bs="x#bs" and f="?f"]
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3851
  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
  3852
    by (simp add: np DVDJ_def)
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3853
  also have "\<dots> = (\<exists> j\<in> {0 .. (n - 1)}. real_of_int n * x = (- ?s) - real_of_int \<lfloor>- ?s\<rfloor> + real_of_int j \<and> real_of_int i rdvd real_of_int (j - \<lfloor>- ?s\<rfloor>))"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3854
    by (simp add: algebra_simps)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3855
  also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3856
  have "\<dots> = (real_of_int i rdvd real_of_int n * x - (-?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
  3857
  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
  3858
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
  3859
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3860
lemma NDVDJ_NDVD:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3861
  assumes xp:"x\<ge> 0" and x1: "x < 1" and np:"real_of_int 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
  3862
  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
  3863
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
  3864
  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
  3865
  let ?s= "Inum (x#bs) s"
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3866
  from foldr_conj_map[where xs="[0..n - 1]" and bs="x#bs" and f="?f"]
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3867
  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
  3868
    by (simp add: np NDVDJ_def)
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  3869
  also have "\<dots> = (\<not> (\<exists> j\<in> {0 .. (n - 1)}. real_of_int n * x = (- ?s) - real_of_int \<lfloor>- ?s\<rfloor> + real_of_int j \<and> real_of_int i rdvd real_of_int (j - \<lfloor>- ?s\<rfloor>)))"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3870
    by (simp add: algebra_simps)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3871
  also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  3872
  have "\<dots> = (\<not> (real_of_int i rdvd real_of_int n * x - (-?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
  3873
  finally show ?thesis by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3874
qed
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3875
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3876
lemma foldr_disj_map_rlfm2:
23264
324622260d29 Added 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
  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
  3878
  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
  3879
using lf by (induct xs, auto)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3880
lemma foldr_And_map_rlfm2:
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3881
  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
  3882
  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
  3883
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
  3884
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3885
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
  3886
  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
  3887
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
  3888
  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
  3889
                         (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
  3890
  have th: "\<forall> j. isrlfm (?f j)" using nb np by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3891
  from DVDJ_def foldr_disj_map_rlfm2[OF th] show ?thesis 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
  3892
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
  3893
324622260d29 Added 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
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
  3895
  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
  3896
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
  3897
  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
  3898
                      (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
  3899
  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
  3900
  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
  3901
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
  3902
23997
a23d0b4b1c1f Updated proofs;
chaieb
parents: 23993
diff changeset
  3903
definition DVD :: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm" where
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3904
  DVD_def: "DVD i c t =
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3905
  (if i=0 then eq c t else
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3906
  if c = 0 then (Dvd i t) else if c >0 then DVDJ \<bar>i\<bar> c t else DVDJ \<bar>i\<bar> (-c) (Neg t))"
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3907
23997
a23d0b4b1c1f Updated proofs;
chaieb
parents: 23993
diff changeset
  3908
definition  NDVD :: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm" where
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3909
  "NDVD i c t =
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3910
  (if i=0 then neq c t else
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3911
  if c = 0 then (NDvd i t) else if c >0 then NDVDJ \<bar>i\<bar> c t else NDVDJ \<bar>i\<bar> (-c) (Neg 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
  3912
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3913
lemma DVD_mono:
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3914
  assumes xp: "0\<le> x" and x1: "x < 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
  3915
  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
  3916
  (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
  3917
proof(clarify)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3918
  fix a n 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
  3919
  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
  3920
  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
  3921
  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
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3922
  moreover {assume iz: "i=0" hence ?th using eq_mono[rule_format, OF conjI[OF H nb]]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3923
      by (simp add: DVD_def rdvd_left_0_eq)}
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3924
  moreover {assume inz: "i\<noteq>0" and "n=0" hence ?th by (simp add: H DVD_def) }
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3925
  moreover {assume inz: "i\<noteq>0" and "n<0" hence ?th
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3926
      by (simp add: DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3927
        rdvd_minus[where d="i" and t="real_of_int 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
  3928
  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
  3929
  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
  3930
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
  3931
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3932
lemma NDVD_mono:   assumes xp: "0\<le> x" and x1: "x < 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
  3933
  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
  3934
  (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
  3935
proof(clarify)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3936
  fix a n 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
  3937
  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
  3938
  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
  3939
  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
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3940
  moreover {assume iz: "i=0" hence ?th using neq_mono[rule_format, OF conjI[OF H nb]]
23264
324622260d29 Added 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
      by (simp add: NDVD_def rdvd_left_0_eq)}
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3942
  moreover {assume inz: "i\<noteq>0" and "n=0" hence ?th by (simp add: H NDVD_def) }
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3943
  moreover {assume inz: "i\<noteq>0" and "n<0" hence ?th
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3944
      by (simp add: NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3945
        rdvd_minus[where d="i" and t="real_of_int n * x + Inum (x # bs) s"]) }
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3946
  moreover {assume inz: "i\<noteq>0" and "n>0" hence ?th
23264
324622260d29 Added 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
      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
  3948
  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
  3949
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
  3950
324622260d29 Added 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
lemma DVD_l: "isrlfm (rsplit (DVD i) a)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3952
  by (rule rsplit_l[where f="DVD i" and a="a"], auto simp add: DVD_def eq_def DVDJ_l)
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  3953
(case_tac s, simp_all, rename_tac nat a b, case_tac "nat", 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
  3954
324622260d29 Added 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
lemma NDVD_l: "isrlfm (rsplit (NDVD i) a)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3956
  by (rule rsplit_l[where f="NDVD i" and a="a"], auto simp add: NDVD_def neq_def NDVDJ_l)
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  3957
(case_tac s, simp_all, rename_tac nat a b, case_tac "nat", 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
  3958
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3959
fun rlfm :: "fm \<Rightarrow> fm"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3960
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
  3961
  "rlfm (And p q) = conj (rlfm p) (rlfm q)"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3962
| "rlfm (Or p q) = disj (rlfm p) (rlfm q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3963
| "rlfm (Imp p q) = disj (rlfm (NOT p)) (rlfm q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3964
| "rlfm (Iff p q) = disj (conj(rlfm p) (rlfm q)) (conj(rlfm (NOT p)) (rlfm (NOT q)))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3965
| "rlfm (Lt a) = rsplit lt a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3966
| "rlfm (Le a) = rsplit le a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3967
| "rlfm (Gt a) = rsplit gt a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3968
| "rlfm (Ge a) = rsplit ge a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3969
| "rlfm (Eq a) = rsplit eq a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3970
| "rlfm (NEq a) = rsplit neq a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3971
| "rlfm (Dvd i a) = rsplit (\<lambda> t. DVD i t) a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3972
| "rlfm (NDvd i a) = rsplit (\<lambda> t. NDVD i t) a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3973
| "rlfm (NOT (And p q)) = disj (rlfm (NOT p)) (rlfm (NOT q))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3974
| "rlfm (NOT (Or p q)) = conj (rlfm (NOT p)) (rlfm (NOT q))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3975
| "rlfm (NOT (Imp p q)) = conj (rlfm p) (rlfm (NOT q))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3976
| "rlfm (NOT (Iff p q)) = disj (conj(rlfm p) (rlfm(NOT q))) (conj(rlfm(NOT p)) (rlfm q))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3977
| "rlfm (NOT (NOT p)) = rlfm p"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3978
| "rlfm (NOT T) = F"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3979
| "rlfm (NOT F) = T"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3980
| "rlfm (NOT (Lt a)) = simpfm (rlfm (Ge a))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3981
| "rlfm (NOT (Le a)) = simpfm (rlfm (Gt a))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3982
| "rlfm (NOT (Gt a)) = simpfm (rlfm (Le a))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3983
| "rlfm (NOT (Ge a)) = simpfm (rlfm (Lt a))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3984
| "rlfm (NOT (Eq a)) = simpfm (rlfm (NEq a))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3985
| "rlfm (NOT (NEq a)) = simpfm (rlfm (Eq a))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3986
| "rlfm (NOT (Dvd i a)) = simpfm (rlfm (NDvd i a))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3987
| "rlfm (NOT (NDvd i a)) = simpfm (rlfm (Dvd i a))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  3988
| "rlfm 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
  3989
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3990
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
  3991
  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
  3992
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3993
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
  3994
proof (induct p)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3995
  case (Lt 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
  3996
  hence "bound0 (Lt a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  3997
    by (cases a,simp_all, rename_tac nat a b, case_tac "nat", 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
  3998
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  3999
  {assume "bound0 (Lt a)" hence bn:"bound0 (simpfm (Lt 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
  4000
      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
  4001
    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
  4002
    with bn bound0at_l have ?case by blast}
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4003
  moreover
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4004
  { 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
  4005
    { 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
  4006
      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
  4007
      have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4008
      from \<open>c > 0\<close> have th:"numgcd (CN 0 c (simpnum e)) \<le> c"
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  4009
        by (simp add: numgcd_def)
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  4010
      from \<open>c > 0\<close> have th': "c\<noteq>0" by auto
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  4011
      from \<open>c > 0\<close> have cp: "c \<ge> 0" by simp
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  4012
      from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4013
      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
  4014
    }
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4015
    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
  4016
      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
  4017
  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
  4018
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4019
  case (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
  4020
  hence "bound0 (Le a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  4021
    by (cases a,simp_all, rename_tac nat a b, case_tac "nat", 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
  4022
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4023
  { 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
  4024
      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
  4025
    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
  4026
    with bn bound0at_l have ?case by blast}
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4027
  moreover
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4028
  { 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
  4029
    { 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
  4030
      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
  4031
      have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4032
      from \<open>c > 0\<close> have th:"numgcd (CN 0 c (simpnum e)) \<le> c"
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  4033
        by (simp add: numgcd_def)
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  4034
      from \<open>c > 0\<close> have th': "c\<noteq>0" by auto
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  4035
      from \<open>c > 0\<close> have cp: "c \<ge> 0" by simp
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  4036
      from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4037
      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
  4038
    }
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4039
    with Le a have ?case
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  4040
      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
  4041
  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
  4042
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4043
  case (Gt 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
  4044
  hence "bound0 (Gt a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  4045
    by (cases a, simp_all, rename_tac nat a b,case_tac "nat", 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
  4046
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4047
  {assume "bound0 (Gt a)" hence bn:"bound0 (simpfm (Gt 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
  4048
      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
  4049
    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
  4050
    with bn bound0at_l have ?case by blast}
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4051
  moreover
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4052
  { 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
  4053
    { 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
  4054
      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
  4055
      have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4056
      from \<open>c > 0\<close> have th:"numgcd (CN 0 c (simpnum e)) \<le> c"
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  4057
        by (simp add: numgcd_def)
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  4058
      from \<open>c > 0\<close> have th': "c\<noteq>0" by auto
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  4059
      from \<open>c > 0\<close> have cp: "c \<ge> 0" by simp
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  4060
      from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4061
      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
  4062
    }
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4063
    with Gt a have ?case
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  4064
      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
  4065
  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
  4066
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4067
  case (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
  4068
  hence "bound0 (Ge a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  4069
    by (cases a,simp_all, rename_tac nat a b, case_tac "nat", 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
  4070
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4071
  { 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
  4072
      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
  4073
    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
  4074
    with bn bound0at_l have ?case by blast}
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4075
  moreover
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4076
  { 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
  4077
    { 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
  4078
      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
  4079
      have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4080
      from \<open>c > 0\<close> have th:"numgcd (CN 0 c (simpnum e)) \<le> c"
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  4081
        by (simp add: numgcd_def)
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  4082
      from \<open>c > 0\<close> have th': "c\<noteq>0" by auto
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  4083
      from \<open>c > 0\<close> have cp: "c \<ge> 0" by simp
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  4084
      from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4085
      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
  4086
    }
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4087
    with Ge a have ?case
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  4088
      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
  4089
  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
  4090
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4091
  case (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
  4092
  hence "bound0 (Eq a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  4093
    by (cases a,simp_all, rename_tac nat a b, case_tac "nat", 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
  4094
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4095
  { 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
  4096
      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
  4097
    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
  4098
    with bn bound0at_l have ?case by blast}
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4099
  moreover
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4100
  { 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
  4101
    { 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
  4102
      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
  4103
      have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4104
      from \<open>c > 0\<close> have th:"numgcd (CN 0 c (simpnum e)) \<le> c"
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  4105
        by (simp add: numgcd_def)
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  4106
      from \<open>c > 0\<close> have th': "c\<noteq>0" by auto
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  4107
      from \<open>c > 0\<close> have cp: "c \<ge> 0" by simp
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  4108
      from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4109
      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
  4110
    }
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4111
    with Eq a have ?case
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  4112
      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
  4113
  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
  4114
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4115
  case (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
  4116
  hence "bound0 (NEq a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  4117
    by (cases a,simp_all, rename_tac nat a b, case_tac "nat", 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
  4118
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4119
  {assume "bound0 (NEq a)" hence bn:"bound0 (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
  4120
      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
  4121
    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
  4122
    with bn bound0at_l have ?case by blast}
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4123
  moreover
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4124
  { 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
  4125
    { 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
  4126
      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
  4127
      have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4128
      from \<open>c > 0\<close> have th:"numgcd (CN 0 c (simpnum e)) \<le> c"
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  4129
        by (simp add: numgcd_def)
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  4130
      from \<open>c > 0\<close> have th': "c\<noteq>0" by auto
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  4131
      from \<open>c > 0\<close> have cp: "c \<ge> 0" by simp
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  4132
      from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4133
      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
  4134
    }
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4135
    with NEq a have ?case
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  4136
      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
  4137
  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
  4138
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4139
  case (Dvd i a) hence "bound0 (Dvd i a)" by auto hence bn:"bound0 (simpfm (Dvd i 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
  4140
    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
  4141
  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
  4142
  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
  4143
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4144
  case (NDvd i a)  hence "bound0 (NDvd i a)" by auto hence bn:"bound0 (simpfm (NDvd i 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
  4145
    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
  4146
  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
  4147
  with bn bound0at_l show ?case by blast
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  4148
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
  4149
324622260d29 Added 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
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
  4151
  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
  4152
  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
  4153
  shows "(Ifm (x#bs) (rlfm p) = Ifm (x# bs) p) \<and> isrlfm (rlfm p)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4154
  using qfp
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4155
by (induct p rule: rlfm.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
  4156
(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
  4157
               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
  4158
               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
  4159
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
  4160
  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
  4161
  shows "isrlfm (rlfm p)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4162
  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
  4163
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
  4164
324622260d29 Added 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
    (* 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
  4166
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
  4167
  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
  4168
  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
  4169
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
  4170
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
  4171
  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
  4172
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
  4173
  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
  4174
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4175
  case (3 c e)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4176
  from 3 have nb: "numbound0 e" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4177
  from 3 have cp: "real_of_int c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4178
  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
  4179
  let ?e="Inum (a#bs) e"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4180
  let ?z = "(- ?e) / real_of_int 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
  4181
  {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
  4182
    assume xz: "x < ?z"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4183
    hence "(real_of_int c * x < - ?e)"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4184
      by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4185
    hence "real_of_int c * x + ?e < 0" by arith
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4186
    hence "real_of_int c * x + ?e \<noteq> 0" 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
  4187
    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
  4188
      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
  4189
  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
  4190
  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
  4191
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4192
  case (4 c e)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4193
  from 4 have nb: "numbound0 e" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4194
  from 4 have cp: "real_of_int c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4195
  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
  4196
  let ?e="Inum (a#bs) e"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4197
  let ?z = "(- ?e) / real_of_int 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
  4198
  {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
  4199
    assume xz: "x < ?z"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4200
    hence "(real_of_int c * x < - ?e)"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4201
      by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4202
    hence "real_of_int c * x + ?e < 0" by arith
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4203
    hence "real_of_int c * x + ?e \<noteq> 0" 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
  4204
    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
  4205
      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
  4206
  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
  4207
  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
  4208
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4209
  case (5 c e)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4210
  from 5 have nb: "numbound0 e" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4211
  from 5 have cp: "real_of_int c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4212
  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
  4213
  let ?e="Inum (a#bs) e"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4214
  let ?z = "(- ?e) / real_of_int 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
  4215
  {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
  4216
    assume xz: "x < ?z"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4217
    hence "(real_of_int c * x < - ?e)"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4218
      by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4219
    hence "real_of_int c * x + ?e < 0" 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
  4220
    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
  4221
      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
  4222
  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
  4223
  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
  4224
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4225
  case (6 c e)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4226
  from 6 have nb: "numbound0 e" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4227
  from 6 have cp: "real_of_int c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4228
  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
  4229
  let ?e="Inum (a#bs) e"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4230
  let ?z = "(- ?e) / real_of_int 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
  4231
  {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
  4232
    assume xz: "x < ?z"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4233
    hence "(real_of_int c * x < - ?e)"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4234
      by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4235
    hence "real_of_int c * x + ?e < 0" 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
  4236
    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
  4237
      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
  4238
  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
  4239
  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
  4240
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4241
  case (7 c e)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4242
  from 7 have nb: "numbound0 e" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4243
  from 7 have cp: "real_of_int c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4244
  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
  4245
  let ?e="Inum (a#bs) e"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4246
  let ?z = "(- ?e) / real_of_int 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
  4247
  {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
  4248
    assume xz: "x < ?z"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4249
    hence "(real_of_int c * x < - ?e)"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4250
      by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4251
    hence "real_of_int c * x + ?e < 0" 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
  4252
    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
  4253
      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
  4254
  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
  4255
  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
  4256
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4257
  case (8 c e)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4258
  from 8 have nb: "numbound0 e" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4259
  from 8 have cp: "real_of_int c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4260
  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
  4261
  let ?e="Inum (a#bs) e"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4262
  let ?z = "(- ?e) / real_of_int 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
  4263
  {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
  4264
    assume xz: "x < ?z"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4265
    hence "(real_of_int c * x < - ?e)"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4266
      by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4267
    hence "real_of_int c * x + ?e < 0" 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
  4268
    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
  4269
      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
  4270
  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
  4271
  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
  4272
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
  4273
324622260d29 Added 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
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
  4275
  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
  4276
  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
  4277
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
  4278
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
  4279
  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
  4280
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
  4281
  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
  4282
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4283
  case (3 c e)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4284
  from 3 have nb: "numbound0 e" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4285
  from 3 have cp: "real_of_int c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4286
  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
  4287
  let ?e="Inum (a#bs) e"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4288
  let ?z = "(- ?e) / real_of_int 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
  4289
  {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
  4290
    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
  4291
    with mult_strict_right_mono [OF xz cp] cp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4292
    have "(real_of_int c * x > - ?e)" by (simp add: ac_simps)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4293
    hence "real_of_int c * x + ?e > 0" by arith
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4294
    hence "real_of_int c * x + ?e \<noteq> 0" 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
  4295
    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
  4296
      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
  4297
  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
  4298
  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
  4299
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4300
  case (4 c e)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4301
  from 4 have nb: "numbound0 e" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4302
  from 4 have cp: "real_of_int c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4303
  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
  4304
  let ?e="Inum (a#bs) e"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4305
  let ?z = "(- ?e) / real_of_int 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
  4306
  {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
  4307
    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
  4308
    with mult_strict_right_mono [OF xz cp] cp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4309
    have "(real_of_int c * x > - ?e)" by (simp add: ac_simps)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4310
    hence "real_of_int c * x + ?e > 0" by arith
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4311
    hence "real_of_int c * x + ?e \<noteq> 0" 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
  4312
    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
  4313
      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
  4314
  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
  4315
  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
  4316
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4317
  case (5 c e)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4318
  from 5 have nb: "numbound0 e" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4319
  from 5 have cp: "real_of_int c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4320
  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
  4321
  let ?e="Inum (a#bs) e"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4322
  let ?z = "(- ?e) / real_of_int 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
  4323
  {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
  4324
    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
  4325
    with mult_strict_right_mono [OF xz cp] cp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4326
    have "(real_of_int c * x > - ?e)" by (simp add: ac_simps)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4327
    hence "real_of_int c * x + ?e > 0" 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
  4328
    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
  4329
      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
  4330
  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
  4331
  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
  4332
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4333
  case (6 c e)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4334
  from 6 have nb: "numbound0 e" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4335
  from 6 have cp: "real_of_int c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4336
  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
  4337
  let ?e="Inum (a#bs) e"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4338
  let ?z = "(- ?e) / real_of_int 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
  4339
  {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
  4340
    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
  4341
    with mult_strict_right_mono [OF xz cp] cp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4342
    have "(real_of_int c * x > - ?e)" by (simp add: ac_simps)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4343
    hence "real_of_int c * x + ?e > 0" 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
  4344
    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
  4345
      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
  4346
  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
  4347
  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
  4348
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4349
  case (7 c e)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4350
  from 7 have nb: "numbound0 e" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4351
  from 7 have cp: "real_of_int c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4352
  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
  4353
  let ?e="Inum (a#bs) e"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4354
  let ?z = "(- ?e) / real_of_int 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
  4355
  {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
  4356
    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
  4357
    with mult_strict_right_mono [OF xz cp] cp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4358
    have "(real_of_int c * x > - ?e)" by (simp add: ac_simps)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4359
    hence "real_of_int c * x + ?e > 0" 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
  4360
    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
  4361
      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
  4362
  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
  4363
  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
  4364
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4365
  case (8 c e)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4366
  from 8 have nb: "numbound0 e" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4367
  from 8 have cp: "real_of_int c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4368
  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
  4369
  let ?e="Inum (a#bs) e"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4370
  let ?z = "(- ?e) / real_of_int 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
  4371
  {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
  4372
    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
  4373
    with mult_strict_right_mono [OF xz cp] cp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4374
    have "(real_of_int c * x > - ?e)" by (simp add: ac_simps)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4375
    hence "real_of_int c * x + ?e > 0" 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
  4376
    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
  4377
      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
  4378
  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
  4379
  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
  4380
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
  4381
324622260d29 Added 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
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
  4383
  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
  4384
  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
  4385
  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
  4386
  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
  4387
324622260d29 Added 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
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
  4389
  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
  4390
  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
  4391
  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
  4392
  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
  4393
324622260d29 Added 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
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
  4395
  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
  4396
  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
  4397
  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
  4398
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
  4399
  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
  4400
  have th: "\<forall> x. Ifm (x#bs) (minusinf p)" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4401
  from rminusinf_inf[OF lp, where 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
  4402
  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
  4403
  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
  4404
  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
  4405
  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
  4406
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
  4407
324622260d29 Added 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
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
  4409
  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
  4410
  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
  4411
  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
  4412
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
  4413
  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
  4414
  have th: "\<forall> x. Ifm (x#bs) (plusinf p)" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4415
  from rplusinf_inf[OF lp, where 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
  4416
  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
  4417
  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
  4418
  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
  4419
  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
  4420
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
  4421
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4422
fun \<Upsilon>:: "fm \<Rightarrow> (num \<times> int) list"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4423
where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4424
  "\<Upsilon> (And p q) = (\<Upsilon> p @ \<Upsilon> q)"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4425
| "\<Upsilon> (Or p q) = (\<Upsilon> p @ \<Upsilon> q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4426
| "\<Upsilon> (Eq  (CN 0 c e)) = [(Neg e,c)]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4427
| "\<Upsilon> (NEq (CN 0 c e)) = [(Neg e,c)]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4428
| "\<Upsilon> (Lt  (CN 0 c e)) = [(Neg e,c)]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4429
| "\<Upsilon> (Le  (CN 0 c e)) = [(Neg e,c)]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4430
| "\<Upsilon> (Gt  (CN 0 c e)) = [(Neg e,c)]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4431
| "\<Upsilon> (Ge  (CN 0 c e)) = [(Neg e,c)]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4432
| "\<Upsilon> p = []"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4433
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4434
fun \<upsilon> :: "fm \<Rightarrow> num \<times> int \<Rightarrow> fm"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4435
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
  4436
  "\<upsilon> (And p q) = (\<lambda> (t,n). And (\<upsilon> p (t,n)) (\<upsilon> q (t,n)))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4437
| "\<upsilon> (Or p q) = (\<lambda> (t,n). Or (\<upsilon> p (t,n)) (\<upsilon> q (t,n)))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4438
| "\<upsilon> (Eq (CN 0 c e)) = (\<lambda> (t,n). Eq (Add (Mul c t) (Mul n e)))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4439
| "\<upsilon> (NEq (CN 0 c e)) = (\<lambda> (t,n). NEq (Add (Mul c t) (Mul n e)))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4440
| "\<upsilon> (Lt (CN 0 c e)) = (\<lambda> (t,n). Lt (Add (Mul c t) (Mul n e)))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4441
| "\<upsilon> (Le (CN 0 c e)) = (\<lambda> (t,n). Le (Add (Mul c t) (Mul n e)))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4442
| "\<upsilon> (Gt (CN 0 c e)) = (\<lambda> (t,n). Gt (Add (Mul c t) (Mul n e)))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4443
| "\<upsilon> (Ge (CN 0 c e)) = (\<lambda> (t,n). Ge (Add (Mul c t) (Mul n e)))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4444
| "\<upsilon> p = (\<lambda> (t,n). 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
  4445
324622260d29 Added 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
lemma \<upsilon>_I: assumes lp: "isrlfm p"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4447
  and np: "real_of_int n > 0" and nbt: "numbound0 t"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4448
  shows "(Ifm (x#bs) (\<upsilon> p (t,n)) = Ifm (((Inum (x#bs) t)/(real_of_int 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> _")
23264
324622260d29 Added 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 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
  4450
proof(induct p rule: \<upsilon>.induct)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4451
  case (5 c e)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4452
  from 5 have cp: "c >0" and nb: "numbound0 e" by simp_all
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4453
  have "?I ?u (Lt (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x e) < 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
  4454
    using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4455
  also have "\<dots> = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) < 0)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4456
    by (simp only: pos_less_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)"
64240
eabf80376aab more standardized names
haftmann
parents: 63600
diff changeset
  4457
      and b="0", simplified div_0]) (simp only: algebra_simps)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4458
  also have "\<dots> = (real_of_int c *?t + ?n* (?N x e) < 0)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4459
    using np by simp
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4460
  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
  4461
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4462
  case (6 c e)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4463
  from 6 have cp: "c >0" and nb: "numbound0 e" by simp_all
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4464
  have "?I ?u (Le (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x e) \<le> 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
  4465
    using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4466
  also have "\<dots> = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) \<le> 0)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4467
    by (simp only: pos_le_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)"
64240
eabf80376aab more standardized names
haftmann
parents: 63600
diff changeset
  4468
      and b="0", simplified div_0]) (simp only: algebra_simps)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4469
  also have "\<dots> = (real_of_int c *?t + ?n* (?N x e) \<le> 0)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4470
    using np by simp
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4471
  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
  4472
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4473
  case (7 c e)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4474
  from 7 have cp: "c >0" and nb: "numbound0 e" by simp_all
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4475
  have "?I ?u (Gt (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x e) > 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
  4476
    using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4477
  also have "\<dots> = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) > 0)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4478
    by (simp only: pos_divide_less_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)"
64240
eabf80376aab more standardized names
haftmann
parents: 63600
diff changeset
  4479
      and b="0", simplified div_0]) (simp only: algebra_simps)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4480
  also have "\<dots> = (real_of_int c *?t + ?n* (?N x e) > 0)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4481
    using np by simp
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4482
  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
  4483
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4484
  case (8 c e)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4485
  from 8 have cp: "c >0" and nb: "numbound0 e" by simp_all
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4486
  have "?I ?u (Ge (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x e) \<ge> 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
  4487
    using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4488
  also have "\<dots> = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) \<ge> 0)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4489
    by (simp only: pos_divide_le_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)"
64240
eabf80376aab more standardized names
haftmann
parents: 63600
diff changeset
  4490
      and b="0", simplified div_0]) (simp only: algebra_simps)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4491
  also have "\<dots> = (real_of_int c *?t + ?n* (?N x e) \<ge> 0)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4492
    using np by simp
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4493
  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
  4494
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4495
  case (3 c e)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4496
  from 3 have cp: "c >0" and nb: "numbound0 e" by simp_all
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4497
  from np have np: "real_of_int n \<noteq> 0" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4498
  have "?I ?u (Eq (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x e) = 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
  4499
    using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4500
  also have "\<dots> = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) = 0)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4501
    by (simp only: nonzero_eq_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)"
64240
eabf80376aab more standardized names
haftmann
parents: 63600
diff changeset
  4502
      and b="0", simplified div_0]) (simp only: algebra_simps)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4503
  also have "\<dots> = (real_of_int c *?t + ?n* (?N x e) = 0)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4504
    using np by simp
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4505
  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
  4506
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4507
  case (4 c e)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4508
  from 4 have cp: "c >0" and nb: "numbound0 e" by simp_all
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4509
  from np have np: "real_of_int n \<noteq> 0" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4510
  have "?I ?u (NEq (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x 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
  4511
    using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4512
  also have "\<dots> = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) \<noteq> 0)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4513
    by (simp only: nonzero_eq_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)"
64240
eabf80376aab more standardized names
haftmann
parents: 63600
diff changeset
  4514
      and b="0", simplified div_0]) (simp only: algebra_simps)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4515
  also have "\<dots> = (real_of_int c *?t + ?n* (?N x e) \<noteq> 0)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4516
    using np by simp
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4517
  finally show ?case using nbt nb by (simp add: algebra_simps)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4518
qed(simp_all add: nbt numbound0_I[where bs ="bs" and b="(Inum (x#bs) t)/ real_of_int 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
  4519
324622260d29 Added 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
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
  4521
  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
  4522
  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
  4523
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
  4524
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
  4525
324622260d29 Added 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
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
  4527
  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
  4528
  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
  4529
  and ex: "Ifm (x#bs) p" (is "?I x p")
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4530
  shows "\<exists> (s,m) \<in> set (\<Upsilon> p). x \<ge> Inum (a#bs) s / real_of_int m" (is "\<exists> (s,m) \<in> ?U p. x \<ge> ?N a s / real_of_int 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
  4531
proof-
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4532
  have "\<exists> (s,m) \<in> set (\<Upsilon> p). real_of_int m * x \<ge> Inum (a#bs) s " (is "\<exists> (s,m) \<in> ?U p. real_of_int m *x \<ge> ?N a 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
  4533
    using lp nmi ex
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  4534
    by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"])
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4535
  then obtain s m where smU: "(s,m) \<in> set (\<Upsilon> p)" and mx: "real_of_int m * x \<ge> ?N a s" by blast
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4536
  from \<Upsilon>_l[OF lp] smU have mp: "real_of_int m > 0" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4537
  from pos_divide_le_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \<ge> ?N a s / real_of_int m"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57492
diff changeset
  4538
    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
  4539
  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
  4540
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
  4541
324622260d29 Added 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
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
  4543
  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
  4544
  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
  4545
  and ex: "Ifm (x#bs) p" (is "?I x p")
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4546
  shows "\<exists> (s,m) \<in> set (\<Upsilon> p). x \<le> Inum (a#bs) s / real_of_int m" (is "\<exists> (s,m) \<in> ?U p. x \<le> ?N a s / real_of_int 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
  4547
proof-
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4548
  have "\<exists> (s,m) \<in> set (\<Upsilon> p). real_of_int m * x \<le> Inum (a#bs) s " (is "\<exists> (s,m) \<in> ?U p. real_of_int m *x \<le> ?N a 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
  4549
    using lp nmi ex
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  4550
    by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"])
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4551
  then obtain s m where smU: "(s,m) \<in> set (\<Upsilon> p)" and mx: "real_of_int m * x \<le> ?N a s" by blast
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4552
  from \<Upsilon>_l[OF lp] smU have mp: "real_of_int m > 0" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4553
  from pos_le_divide_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \<le> ?N a s / real_of_int m"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57492
diff changeset
  4554
    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
  4555
  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
  4556
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
  4557
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4558
lemma lin_dense:
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4559
  assumes lp: "isrlfm p"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4560
  and noS: "\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> (\<lambda> (t,n). Inum (x#bs) t / real_of_int n) ` set (\<Upsilon> p)"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4561
  (is "\<forall> t. _ \<and> _ \<longrightarrow> t \<notin> (\<lambda> (t,n). ?N x t / real_of_int n ) ` (?U 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
  4562
  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
  4563
  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
  4564
  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
  4565
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
  4566
proof (induct p rule: isrlfm.induct)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4567
  case (5 c e) hence cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4568
  from 5 have "x * real_of_int c + ?N x e < 0" by (simp add: algebra_simps)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4569
  hence pxc: "x < (- ?N x e) / real_of_int c"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4570
    by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="-?N x e"])
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4571
  from 5 have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real_of_int c" by auto
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4572
  with ly yu have yne: "y \<noteq> - ?N x e / real_of_int c" by auto
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4573
  hence "y < (- ?N x e) / real_of_int c \<or> y > (-?N x e) / real_of_int c" by auto
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4574
  moreover {assume y: "y < (-?N x e)/ real_of_int c"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4575
    hence "y * real_of_int c < - ?N x e"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4576
      by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric])
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4577
    hence "real_of_int c * y + ?N x e < 0" by (simp add: algebra_simps)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4578
    hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4579
  moreover {assume y: "y > (- ?N x e) / real_of_int c"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4580
    with yu have eu: "u > (- ?N x e) / real_of_int c" by auto
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4581
    with noSc ly yu have "(- ?N x e) / real_of_int c \<le> l" by (cases "(- ?N x e) / real_of_int c > l", auto)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4582
    with lx pxc have "False" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4583
    hence ?case by simp }
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4584
  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
  4585
next
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4586
  case (6 c e) hence cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4587
  from 6 have "x * real_of_int c + ?N x e \<le> 0" by (simp add: algebra_simps)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4588
  hence pxc: "x \<le> (- ?N x e) / real_of_int c"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4589
    by (simp only: pos_le_divide_eq[OF cp, where a="x" and b="-?N x e"])
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4590
  from 6 have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real_of_int c" by auto
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4591
  with ly yu have yne: "y \<noteq> - ?N x e / real_of_int c" by auto
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4592
  hence "y < (- ?N x e) / real_of_int c \<or> y > (-?N x e) / real_of_int c" by auto
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4593
  moreover {assume y: "y < (-?N x e)/ real_of_int c"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4594
    hence "y * real_of_int c < - ?N x e"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4595
      by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric])
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4596
    hence "real_of_int c * y + ?N x e < 0" by (simp add: algebra_simps)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4597
    hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4598
  moreover {assume y: "y > (- ?N x e) / real_of_int c"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4599
    with yu have eu: "u > (- ?N x e) / real_of_int c" by auto
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4600
    with noSc ly yu have "(- ?N x e) / real_of_int c \<le> l" by (cases "(- ?N x e) / real_of_int c > l", auto)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4601
    with lx pxc have "False" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4602
    hence ?case by simp }
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4603
  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
  4604
next
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4605
  case (7 c e) hence cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4606
  from 7 have "x * real_of_int c + ?N x e > 0" by (simp add: algebra_simps)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4607
  hence pxc: "x > (- ?N x e) / real_of_int c"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4608
    by (simp only: pos_divide_less_eq[OF cp, where a="x" and b="-?N x e"])
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4609
  from 7 have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real_of_int c" by auto
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4610
  with ly yu have yne: "y \<noteq> - ?N x e / real_of_int c" by auto
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4611
  hence "y < (- ?N x e) / real_of_int c \<or> y > (-?N x e) / real_of_int c" by auto
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4612
  moreover {assume y: "y > (-?N x e)/ real_of_int c"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4613
    hence "y * real_of_int c > - ?N x e"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4614
      by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric])
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4615
    hence "real_of_int c * y + ?N x e > 0" by (simp add: algebra_simps)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4616
    hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4617
  moreover {assume y: "y < (- ?N x e) / real_of_int c"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4618
    with ly have eu: "l < (- ?N x e) / real_of_int c" by auto
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4619
    with noSc ly yu have "(- ?N x e) / real_of_int c \<ge> u" by (cases "(- ?N x e) / real_of_int c > l", auto)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4620
    with xu pxc have "False" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4621
    hence ?case by simp }
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4622
  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
  4623
next
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4624
  case (8 c e) hence cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4625
  from 8 have "x * real_of_int c + ?N x e \<ge> 0" by (simp add: algebra_simps)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4626
  hence pxc: "x \<ge> (- ?N x e) / real_of_int c"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4627
    by (simp only: pos_divide_le_eq[OF cp, where a="x" and b="-?N x e"])
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4628
  from 8 have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real_of_int c" by auto
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4629
  with ly yu have yne: "y \<noteq> - ?N x e / real_of_int c" by auto
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4630
  hence "y < (- ?N x e) / real_of_int c \<or> y > (-?N x e) / real_of_int c" by auto
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4631
  moreover {assume y: "y > (-?N x e)/ real_of_int c"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4632
    hence "y * real_of_int c > - ?N x e"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4633
      by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric])
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4634
    hence "real_of_int c * y + ?N x e > 0" by (simp add: algebra_simps)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4635
    hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4636
  moreover {assume y: "y < (- ?N x e) / real_of_int c"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4637
    with ly have eu: "l < (- ?N x e) / real_of_int c" by auto
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4638
    with noSc ly yu have "(- ?N x e) / real_of_int c \<ge> u" by (cases "(- ?N x e) / real_of_int c > l", auto)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4639
    with xu pxc have "False" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4640
    hence ?case by simp }
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4641
  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
  4642
next
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4643
  case (3 c e) hence cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4644
  from cp have cnz: "real_of_int c \<noteq> 0" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4645
  from 3 have "x * real_of_int c + ?N x e = 0" by (simp add: algebra_simps)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4646
  hence pxc: "x = (- ?N x e) / real_of_int c"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4647
    by (simp only: nonzero_eq_divide_eq[OF cnz, where a="x" and b="-?N x e"])
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4648
  from 3 have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real_of_int c" by auto
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4649
  with lx xu have yne: "x \<noteq> - ?N x e / real_of_int c" by auto
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4650
  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
  4651
next
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4652
  case (4 c e) hence cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4653
  from cp have cnz: "real_of_int c \<noteq> 0" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4654
  from 4 have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real_of_int c" by auto
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4655
  with ly yu have yne: "y \<noteq> - ?N x e / real_of_int c" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4656
  hence "y* real_of_int c \<noteq> -?N x e"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4657
    by (simp only: nonzero_eq_divide_eq[OF cnz, where a="y" and b="-?N x e"]) simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4658
  hence "y* real_of_int c + ?N x e \<noteq> 0" by (simp add: algebra_simps)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4659
  thus ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4660
    by (simp add: algebra_simps)
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  4661
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
  4662
324622260d29 Added 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
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
  4664
  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
  4665
  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
  4666
  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
  4667
  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
  4668
  shows "\<exists> (l,n) \<in> set (\<Upsilon> p). \<exists> (s,m) \<in> set (\<Upsilon> p).
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4669
    ?I ((Inum (x#bs) l / real_of_int n + Inum (x#bs) s / real_of_int 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
  4670
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
  4671
  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
  4672
  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
  4673
  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
  4674
  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
  4675
  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
  4676
  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
  4677
  have npi': "\<not> (?I a (?P p))" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4678
  have "\<exists> (l,n) \<in> set (\<Upsilon> p). \<exists> (s,m) \<in> set (\<Upsilon> p). ?I ((?N a l/real_of_int n + ?N a s /real_of_int 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
  4679
  proof-
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4680
    let ?M = "(\<lambda> (t,c). ?N a t / real_of_int c) ` ?U"
23264
324622260d29 Added 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
    have fM: "finite ?M" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4682
    from rminusinf_\<Upsilon>[OF lp nmi pa] rplusinf_\<Upsilon>[OF lp npi pa]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4683
    have "\<exists> (l,n) \<in> set (\<Upsilon> p). \<exists> (s,m) \<in> set (\<Upsilon> p). a \<le> ?N x l / real_of_int n \<and> a \<ge> ?N x s / real_of_int m" by blast
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4684
    then obtain "t" "n" "s" "m" where
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4685
      tnU: "(t,n) \<in> ?U" and smU: "(s,m) \<in> ?U"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4686
      and xs1: "a \<le> ?N x s / real_of_int m" and tx1: "a \<ge> ?N x t / real_of_int n" by blast
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4687
    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_of_int m" and tx: "a \<ge> ?N a t / real_of_int 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
  4688
    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
  4689
    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
  4690
    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
  4691
    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
  4692
    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
  4693
    have uinM: "?u \<in> ?M" using fM Mne by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4694
    have tnM: "?N a t / real_of_int n \<in> ?M" using tnU by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4695
    have smM: "?N a s / real_of_int m \<in> ?M" using smU 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
  4696
    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
  4697
    have Mu: "\<forall> t\<in> ?M. t \<le> ?u" using Mne fM by auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4698
    have "?l \<le> ?N a t / real_of_int n" using tnM Mne by simp hence lx: "?l \<le> a" using tx by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4699
    have "?N a s / real_of_int m \<le> ?u" using smM Mne by simp hence xu: "a \<le> ?u" using xs 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
  4700
    from finite_set_intervals2[where P="\<lambda> x. ?I x p",OF pa lx xu linM uinM fM lM Mu]
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4701
    have "(\<exists> s\<in> ?M. ?I s p) \<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
  4702
      (\<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
  4703
    moreover { fix u assume um: "u\<in> ?M" and pu: "?I u p"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4704
      hence "\<exists> (tu,nu) \<in> ?U. u = ?N a tu / real_of_int nu" by auto
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4705
      then obtain "tu" "nu" where tuU: "(tu,nu) \<in> ?U" and tuu:"u= ?N a tu / real_of_int nu" by blast
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4706
      have "(u + u) / 2 = u" by auto with pu tuu
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4707
      have "?I (((?N a tu / real_of_int nu) + (?N a tu / real_of_int nu)) / 2) 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
  4708
      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
  4709
    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
  4710
      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"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4711
      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
  4712
        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
  4713
        by blast
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4714
      from t1M have "\<exists> (t1u,t1n) \<in> ?U. t1 = ?N a t1u / real_of_int t1n" by auto
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4715
      then obtain "t1u" "t1n" where t1uU: "(t1u,t1n) \<in> ?U" and t1u: "t1 = ?N a t1u / real_of_int t1n" by blast
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4716
      from t2M have "\<exists> (t2u,t2n) \<in> ?U. t2 = ?N a t2u / real_of_int t2n" by auto
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4717
      then obtain "t2u" "t2n" where t2uU: "(t2u,t2n) \<in> ?U" and t2u: "t2 = ?N a t2u / real_of_int t2n" 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
  4718
      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
  4719
      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
  4720
      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
  4721
      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
  4722
      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
  4723
    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
  4724
  qed
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4725
  then obtain "l" "n" "s"  "m" where lnU: "(l,n) \<in> ?U" and smU:"(s,m) \<in> ?U"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4726
    and pu: "?I ((?N a l / real_of_int n + ?N a s / real_of_int m) / 2) p" 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
  4727
  from lnU smU \<Upsilon>_l[OF lp] have nbl: "numbound0 l" and nbs: "numbound0 s" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4728
  from numbound0_I[OF nbl, 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
  4729
    numbound0_I[OF nbs, where bs="bs" and b="a" and b'="x"] pu
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4730
  have "?I ((?N x l / real_of_int n + ?N x s / real_of_int m) / 2) 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
  4731
  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
  4732
  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
  4733
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
  4734
    (* 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
  4735
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4736
theorem fr_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
  4737
  assumes lp: "isrlfm p"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4738
  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_of_int n + (Inum (x#bs) s) / real_of_int m) /2)#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
  4739
  (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
  4740
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
  4741
  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
  4742
  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
  4743
  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
  4744
  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
  4745
    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
  4746
  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
  4747
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4748
  assume "?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
  4749
  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
  4750
  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
  4751
  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
  4752
  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
  4753
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
  4754
324622260d29 Added 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
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4756
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
  4757
  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
  4758
  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
  4759
  (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
  4760
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
  4761
  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
  4762
  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
  4763
  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
  4764
  moreover {assume nmi: "\<not> ?M" and npi: "\<not> ?P"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4765
    let ?f ="\<lambda> (t,n). Inum (x#bs) t / real_of_int 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
  4766
    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
  4767
    {fix t n s m assume "(t,n)\<in> set (\<Upsilon> p)" and "(s,m) \<in> set (\<Upsilon> p)"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4768
      with \<Upsilon>_l[OF lp] have tnb: "numbound0 t" and np:"real_of_int n > 0" and snb: "numbound0 s" and mp:"real_of_int m > 0"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  4769
        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
  4770
      let ?st = "Add (Mul m t) (Mul n s)"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4771
      from np mp have mnp: "real_of_int (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
  4772
      from tnb snb have st_nb: "numbound0 ?st" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4773
      have st: "(?N t / real_of_int n + ?N s / real_of_int m)/2 = ?N ?st / real_of_int (2*n*m)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  4774
        using mnp mp np by (simp add: algebra_simps add_divide_distrib)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4775
      from \<upsilon>_I[OF lp mnp st_nb, where x="x" and bs="bs"]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4776
      have "?I x (\<upsilon> p (?st,2*n*m)) = ?I ((?N t / real_of_int n + ?N s / real_of_int m) /2) p" by (simp only: st[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
  4777
    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
  4778
  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
  4779
next
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4780
  assume "?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
  4781
  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
  4782
  moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . }
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4783
  moreover {fix t k s l assume "(t,k) \<in> set (\<Upsilon> p)" and "(s,l) \<in> set (\<Upsilon> 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
  4784
    and px:"?I x (\<upsilon> p (Add (Mul l t) (Mul k s), 2*k*l))"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4785
    with \<Upsilon>_l[OF lp] have tnb: "numbound0 t" and np:"real_of_int k > 0" and snb: "numbound0 s" and mp:"real_of_int l > 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
  4786
    let ?st = "Add (Mul l t) (Mul k s)"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4787
    from np mp have mnp: "real_of_int (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
  4788
    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
  4789
    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
  4790
  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
  4791
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
  4792
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  4793
text\<open>The overall Part\<close>
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4794
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4795
lemma real_ex_int_real01:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4796
  shows "(\<exists> (x::real). P x) = (\<exists> (i::int) (u::real). 0\<le> u \<and> u< 1 \<and> P (real_of_int i + u))"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4797
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
  4798
  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
  4799
  assume Px: "P x"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61762
diff changeset
  4800
  let ?i = "\<lfloor>x\<rfloor>"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4801
  let ?u = "x - real_of_int ?i"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4802
  have "x = real_of_int ?i + ?u" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4803
  hence "P (real_of_int ?i + ?u)" using Px by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4804
  moreover have "real_of_int ?i \<le> x" using of_int_floor_le by simp hence "0 \<le> ?u" by arith
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4805
  moreover have "?u < 1" using real_of_int_floor_add_one_gt[where r="x"] by arith
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4806
  ultimately show "(\<exists> (i::int) (u::real). 0\<le> u \<and> u< 1 \<and> P (real_of_int i + u))" 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
  4807
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
  4808
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4809
fun exsplitnum :: "num \<Rightarrow> num"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4810
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
  4811
  "exsplitnum (C c) = (C c)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4812
| "exsplitnum (Bound 0) = Add (Bound 0) (Bound 1)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4813
| "exsplitnum (Bound n) = Bound (n+1)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4814
| "exsplitnum (Neg a) = Neg (exsplitnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4815
| "exsplitnum (Add a b) = Add (exsplitnum a) (exsplitnum b) "
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4816
| "exsplitnum (Sub a b) = Sub (exsplitnum a) (exsplitnum b) "
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4817
| "exsplitnum (Mul c a) = Mul c (exsplitnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4818
| "exsplitnum (Floor a) = Floor (exsplitnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4819
| "exsplitnum (CN 0 c a) = CN 0 c (Add (Mul c (Bound 1)) (exsplitnum a))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4820
| "exsplitnum (CN n c a) = CN (n+1) c (exsplitnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4821
| "exsplitnum (CF c s t) = CF c (exsplitnum s) (exsplitnum t)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4822
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4823
fun exsplit :: "fm \<Rightarrow> fm"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66515
diff changeset
  4824
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
  4825
  "exsplit (Lt a) = Lt (exsplitnum a)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4826
| "exsplit (Le a) = Le (exsplitnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4827
| "exsplit (Gt a) = Gt (exsplitnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4828
| "exsplit (Ge a) = Ge (exsplitnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4829
| "exsplit (Eq a) = Eq (exsplitnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4830
| "exsplit (NEq a) = NEq (exsplitnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4831
| "exsplit (Dvd i a) = Dvd i (exsplitnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4832
| "exsplit (NDvd i a) = NDvd i (exsplitnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4833
| "exsplit (And p q) = And (exsplit p) (exsplit q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4834
| "exsplit (Or p q) = Or (exsplit p) (exsplit q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4835
| "exsplit (Imp p q) = Imp (exsplit p) (exsplit q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4836
| "exsplit (Iff p q) = Iff (exsplit p) (exsplit q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4837
| "exsplit (NOT p) = NOT (exsplit p)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4838
| "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
  4839
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4840
lemma exsplitnum:
23264
324622260d29 Added 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
  "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
  4842
  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
  4843
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4844
lemma 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
  4845
  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
  4846
  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
  4847
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
  4848
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
  4849
324622260d29 Added 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
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
  4851
  assumes qf: "qfree p"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4852
  shows "(Ifm bs (E p)) = (\<exists> (i::int). Ifm (real_of_int i#bs) (E (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (exsplit p))))" (is "?lhs = ?rhs")
23264
324622260d29 Added 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-
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4854
  have "?rhs = (\<exists> (i::int). \<exists> x. 0\<le> x \<and> x < 1 \<and> Ifm (x#(real_of_int i)#bs) (exsplit p))"
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  4855
    by auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4856
  also have "\<dots> = (\<exists> (i::int). \<exists> x. 0\<le> x \<and> x < 1 \<and> Ifm ((real_of_int i + x) #bs) p)"
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  4857
    by (simp only: exsplit[OF qf] ac_simps)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4858
  also have "\<dots> = (\<exists> x. Ifm (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
  4859
    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
  4860
  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
  4861
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
  4862
324622260d29 Added 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
    (* 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
  4864
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  4865
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
  4866
  "ferrack01 p \<equiv> (let p' = rlfm(And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p);
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4867
                    U = remdups(map simp_num_pair
23264
324622260d29 Added 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
                     (map (\<lambda> ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m))
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4869
                           (alluopairs (\<Upsilon> 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
  4870
  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
  4871
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4872
lemma fr_eq_01:
23264
324622260d29 Added 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
  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
  4874
  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
  4875
  (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
  4876
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
  4877
  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
  4878
  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
  4879
  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
  4880
  have MF: "?M = False"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
  4881
    apply (simp add: Let_def reducecoeff_def numgcd_def rsplit_def ge_def lt_def conj_def disj_def)
58410
6d46ad54a2ab explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents: 58310
diff changeset
  4882
    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
  4883
  have PF: "?P = False" apply (simp add: Let_def reducecoeff_def numgcd_def rsplit_def ge_def lt_def conj_def disj_def)
58410
6d46ad54a2ab explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents: 58310
diff changeset
  4884
    by (cases "rlfm p = And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))", simp_all)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4885
  have "(\<exists> x. ?I x ?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
  4886
    ((?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
  4887
    (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
  4888
  proof
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4889
    assume "\<exists> x. ?I x ?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
  4890
    then obtain x where qx: "?I x ?q" by blast
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4891
    hence xp: "0\<le> x" and x1: "x< 1" and px: "?I x 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
  4892
      by (auto simp add: rsplit_def lt_def ge_def rlfm_I[OF qf])
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4893
    from qx have "?I x ?rq "
23264
324622260d29 Added 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
      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
  4895
    hence lqx: "?I x ?rq " using simpfm[where p="?rq" and bs="x#bs"] by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4896
    from qf have qfq:"isrlfm ?rq"
23264
324622260d29 Added 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
      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
  4898
    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
  4899
  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
  4900
    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
  4901
    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
  4902
    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
  4903
    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
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4904
    from qf have lrq:"isrlfm ?rq"using rlfm_l[OF qf]
23264
324622260d29 Added 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
      by (auto simp add: rsplit_def lt_def ge_def)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4906
    from aU bU \<Upsilon>_l[OF lrq] have tnb: "numbound0 t" and np:"real_of_int n > 0" and snb: "numbound0 s" and mp:"real_of_int m > 0" by (auto 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
  4907
    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
  4908
    from tnb snb have stnb: "numbound0 ?st" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4909
    from np mp have mnp: "real_of_int (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
  4910
    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
  4911
    have "\<exists> x. ?I x ?rq" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4912
    thus "?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
  4913
      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
  4914
  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
  4915
  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
  4916
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
  4917
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4918
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
  4919
  assumes Ul: "\<forall> (t,n) \<in> set U. numbound0 t \<and> n >0"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4920
  shows "((\<lambda> (t,n). Inum (x#bs) t /real_of_int 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_of_int n + Inum (x#bs) s /real_of_int m)/2) ` (set U \<times> set U))"
23264
324622260d29 Added 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
  (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
  4922
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
  4923
  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
  4924
  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
  4925
  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
  4926
    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
  4927
  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
  4928
  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
  4929
  from Ul th have mnz: "m \<noteq> 0" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4930
  from Ul th have  nnz: "n \<noteq> 0" by auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4931
  have st: "(?N t / real_of_int n + ?N s / real_of_int m)/2 = ?N ?st / real_of_int (2*n*m)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4932
   using mnz nnz by (simp add: algebra_simps add_divide_distrib)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4933
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4934
  thus "(real_of_int m *  Inum (x # bs) t + real_of_int n * Inum (x # bs) s) /
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4935
       (2 * real_of_int n * real_of_int 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
  4936
       \<in> (\<lambda>((t, n), s, m).
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4937
             (Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m) / 2) `
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4938
         (set U \<times> set U)"using mnz nnz th
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4939
    apply (auto simp add: th add_divide_distrib algebra_simps split_def image_def)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4940
    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
  4941
  (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
  4942
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
  4943
  fix t n s m
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4944
  assume tnU: "(t,n) \<in> set U" and smU:"(s,m) \<in> set U"
23264
324622260d29 Added 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
  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
  4946
  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
  4947
  from Ul smU have mnz: "m \<noteq> 0" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4948
  from Ul tnU have  nnz: "n \<noteq> 0" by auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4949
  have st: "(?N t / real_of_int n + ?N s / real_of_int m)/2 = ?N ?st / real_of_int (2*n*m)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4950
   using mnz nnz by (simp add: algebra_simps add_divide_distrib)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4951
 let ?P = "\<lambda> (t',n') (s',m'). (Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m)/2 = (Inum (x # bs) t' / real_of_int n' + Inum (x # bs) s' / real_of_int m')/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
  4952
 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
  4953
   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
  4954
 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
  4955
 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
  4956
 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
  4957
   by blast
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4958
 then obtain t' n' s' m' where ts'_U: "((t',n'),(s',m')) \<in> set (alluopairs U)"
23264
324622260d29 Added 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
   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
  4960
 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
  4961
 let ?st' = "Add (Mul m' t') (Mul n' s')"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4962
   have st': "(?N t' / real_of_int n' + ?N s' / real_of_int m')/2 = ?N ?st' / real_of_int (2*n'*m')"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4963
   using mnz' nnz' by (simp add: algebra_simps add_divide_distrib)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4964
 from Pts' have
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4965
   "(Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m)/2 = (Inum (x # bs) t' / real_of_int n' + Inum (x # bs) s' / real_of_int m')/2" by simp
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4966
 also have "\<dots> = ((\<lambda>(t, n). Inum (x # bs) t / real_of_int n) ((\<lambda>((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ((t',n'),(s',m'))))" by (simp add: st')
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4967
 finally show "(Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m) / 2
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4968
          \<in> (\<lambda>(t, n). Inum (x # bs) t / real_of_int 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
  4969
            (\<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
  4970
            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
  4971
   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
  4972
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
  4973
324622260d29 Added 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
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
  4975
  assumes lp: "isrlfm p"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4976
  and UU': "((\<lambda> (t,n). Inum (x#bs) t /real_of_int n) ` U') = ((\<lambda> ((t,n),(s,m)). (Inum (x#bs) t /real_of_int n + Inum (x#bs) s /real_of_int m)/2) ` (U \<times> U))" (is "?f ` U' = ?g ` (U\<times>U)")
23264
324622260d29 Added 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
  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
  4978
  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
  4979
  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
  4980
  (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
  4981
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
  4982
  assume ?lhs
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4983
  then obtain t n s m where tnU: "(t,n) \<in> U" and smU:"(s,m) \<in> U" 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
  4984
    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
  4985
  let ?N = "\<lambda> t. Inum (x#bs) t"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4986
  from tnU smU U have tnb: "numbound0 t" and np: "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
  4987
    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
  4988
  let ?st= "Add (Mul m t) (Mul n s)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4989
  from np mp have mnp: "real_of_int (2*n*m) > 0"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4990
      by (simp add: mult.commute of_int_mult[symmetric] del: 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
  4991
    from tnb snb have stnb: "numbound0 ?st" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4992
  have st: "(?N t / real_of_int n + ?N s / real_of_int m)/2 = ?N ?st / real_of_int (2*n*m)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4993
   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
  4994
  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
  4995
  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
  4996
    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
  4997
  then obtain t' n' where tnU': "(t',n') \<in> U'" and th: "?g ((t,n),(s,m)) = ?f (t',n')" by blast
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  4998
  from U' tnU' have tnb': "numbound0 t'" and np': "real_of_int n' > 0" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  4999
  from \<upsilon>_I[OF lp mnp stnb, where bs="bs" and x="x"] Pst
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5000
  have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real_of_int (2 * n * m) # bs) 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
  5001
  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]]
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5002
  have "Ifm (x # bs) (\<upsilon> p (t', n')) " by (simp only: st)
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5003
  then show ?rhs using tnU' 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
  5004
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
  5005
  assume ?rhs
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5006
  then obtain t' n' where tnU': "(t',n') \<in> U'" and Pt': "Ifm (x # bs) (\<upsilon> p (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
  5007
    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
  5008
  from tnU' UU' have "?f (t',n') \<in> ?g ` (U\<times>U)" by blast
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5009
  hence "\<exists> ((t,n),(s,m)) \<in> (U\<times>U). ?f (t',n') = ?g ((t,n),(s,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
  5010
    by auto (rule_tac x="(a,b)" in bexI, auto)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5011
  then obtain t n s m where tnU: "(t,n) \<in> U" and smU:"(s,m) \<in> U" 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
  5012
    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
  5013
    let ?N = "\<lambda> t. Inum (x#bs) t"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5014
  from tnU smU U have tnb: "numbound0 t" and np: "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
  5015
    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
  5016
  let ?st= "Add (Mul m t) (Mul n s)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5017
  from np mp have mnp: "real_of_int (2*n*m) > 0"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5018
      by (simp add: mult.commute of_int_mult[symmetric] del: 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
  5019
    from tnb snb have stnb: "numbound0 ?st" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5020
  have st: "(?N t / real_of_int n + ?N s / real_of_int m)/2 = ?N ?st / real_of_int (2*n*m)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  5021
   using mp np by (simp add: algebra_simps add_divide_distrib)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5022
  from U' tnU' have tnb': "numbound0 t'" and np': "real_of_int n' > 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
  5023
  from \<upsilon>_I[OF lp np' tnb', where bs="bs" and x="x",simplified th[simplified split_def fst_conv snd_conv] st] Pt'
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5024
  have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real_of_int (2 * n * m) # bs) 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
  5025
  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
  5026
qed
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5027
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5028
lemma ferrack01:
23264
324622260d29 Added 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
  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
  5030
  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
  5031
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
  5032
  let ?I = "\<lambda> x p. Ifm (x#bs) p"
26935
ee6bcb1b8953 avoid undeclared variables within proofs;
wenzelm
parents: 26932
diff changeset
  5033
  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
  5034
  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
  5035
  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
  5036
  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
  5037
  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
  5038
  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
  5039
  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
  5040
  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
  5041
  let ?Y = "remdups ?SS"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5042
  let ?f= "(\<lambda> (t,n). ?N t / real_of_int n)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5043
  let ?h = "\<lambda> ((t,n),(s,m)). (?N t/real_of_int n + ?N s/ real_of_int m) /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
  5044
  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
  5045
  let ?ep = "evaldjf (\<upsilon> ?q) ?Y"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5046
  from rlfm_l[OF qf] have lq: "isrlfm ?q"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
  5047
    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
  5048
  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
  5049
  from \<Upsilon>_l[OF lq] have U_l: "\<forall> (t,n) \<in> set ?U. numbound0 t \<and> n > 0" .
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5050
  from U_l UpU
50241
8b0fdeeefef7 eliminated some improper identifiers;
wenzelm
parents: 49962
diff changeset
  5051
  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
  5052
  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
  5053
    by (auto)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5054
  have Y_l: "\<forall> (t,n) \<in> set ?Y. numbound0 t \<and> 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
  5055
  proof-
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5056
    { fix t n assume tnY: "(t,n) \<in> set ?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
  5057
      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
  5058
      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
  5059
        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
  5060
           (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
  5061
      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
  5062
      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
  5063
      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
  5064
      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
  5065
    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
  5066
  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
  5067
324622260d29 Added 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
  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
  5069
  proof-
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5070
     from simp_num_pair_ci[where bs="x#bs"] have
23264
324622260d29 Added 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
    "\<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
  5072
     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
  5073
    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
  5074
    also have "\<dots> = (?f ` set ?S)" by (simp add: th)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5075
    also have "\<dots> = ((?f o ?g) ` set ?Up)"
56154
f0a927235162 more complete set of lemmas wrt. image and composition
haftmann
parents: 55584
diff changeset
  5076
      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
  5077
    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
  5078
      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
  5079
    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
  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
  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
  5082
  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
  5083
    { fix t n assume tnY: "(t,n) \<in> set ?Y"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5084
      with Y_l have tnb: "numbound0 t" and np: "real_of_int n > 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
  5085
      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
  5086
    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
  5087
    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
  5088
  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
  5089
  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
  5090
    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
  5091
324622260d29 Added 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
  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
  5093
    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
  5094
  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]
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5095
    by (simp only: split_def fst_conv snd_conv)
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5096
  also have "\<dots> = (Ifm (x#bs) ?ep)"
23264
324622260d29 Added 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
    using evaldjf_ex[where ps="?Y" and bs = "x#bs" and f="\<upsilon> ?q",symmetric]
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61076
diff changeset
  5098
    by (simp only: split_def prod.collapse)
23264
324622260d29 Added 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
  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
  5100
  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
  5101
  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
  5102
  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
  5103
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
  5104
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5105
lemma cp_thm':
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5106
  assumes lp: "iszlfm p (real_of_int (i::int)#bs)"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5107
  and up: "d_\<beta> p 1" and dd: "d_\<delta> p d" and dp: "d > 0"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5108
  shows "(\<exists> (x::int). Ifm (real_of_int x#bs) p) = ((\<exists> j\<in> {1 .. d}. Ifm (real_of_int j#bs) (minusinf p)) \<or> (\<exists> j\<in> {1.. d}. \<exists> b\<in> (Inum (real_of_int i#bs)) ` set (\<beta> p). Ifm ((b+real_of_int j)#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
  5109
  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
  5110
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  5111
definition unit :: "fm \<Rightarrow> fm \<times> num list \<times> int" where
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5112
  "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
  5113
             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
  5114
             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
  5115
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5116
lemma unit: assumes qf: "qfree p"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5117
  shows "\<And> q B d. unit p = (q,B,d) \<Longrightarrow>
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5118
      ((\<exists> (x::int). Ifm (real_of_int x#bs) p) =
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5119
       (\<exists> (x::int). Ifm (real_of_int x#bs) q)) \<and>
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5120
       (Inum (real_of_int i#bs)) ` set B = (Inum (real_of_int i#bs)) ` set (\<beta> q) \<and>
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5121
       d_\<beta> q 1 \<and> d_\<delta> q d \<and> d >0 \<and> iszlfm q (real_of_int (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
  5122
proof-
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5123
  fix q B 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
  5124
  assume qBd: "unit p = (q,B,d)"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5125
  let ?thes = "((\<exists> (x::int). Ifm (real_of_int x#bs) p) = (\<exists> (x::int). Ifm (real_of_int x#bs) q)) \<and>
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5126
    Inum (real_of_int i#bs) ` set B = Inum (real_of_int i#bs) ` set (\<beta> q) \<and>
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5127
    d_\<beta> q 1 \<and> d_\<delta> q d \<and> 0 < d \<and> iszlfm q (real_of_int i # bs) \<and> (\<forall> b\<in> set B. numbound0 b)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5128
  let ?I = "\<lambda> (x::int) p. Ifm (real_of_int 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
  5129
  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
  5130
  let ?l = "\<zeta> ?p'"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5131
  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
  5132
  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
  5133
  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
  5134
  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
  5135
  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
  5136
  let ?A'= "remdups (map simpnum (\<alpha> ?q))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5137
  from conjunct1[OF zlfm_I[OF qf, where 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
  5138
  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
  5139
  from iszlfm_gen[OF conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]]]
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5140
  have lp': "\<forall> (i::int). iszlfm ?p' (real_of_int i#bs)" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5141
  hence lp'': "iszlfm ?p' (real_of_int (i::int)#bs)" by simp
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5142
  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
  5143
  from a_\<beta>_ex[where p="?p'" and l="?l" and bs="bs", OF lp'' dl lp] pp'
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5144
  have pq_ex:"(\<exists> (x::int). ?I x p) = (\<exists> x. ?I x ?q)" by (simp add: int_rdvd_iff)
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5145
  from lp'' lp a_\<beta>[OF lp'' dl lp] have lq:"iszlfm ?q (real_of_int 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
  5146
    by (auto simp add: isint_def)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5147
  from \<delta>[OF lq] have dp:"?d >0" and dd: "d_\<delta> ?q ?d" by blast+
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5148
  let ?N = "\<lambda> t. Inum (real_of_int (i::int)#bs) t"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5149
  have "?N ` set ?B' = ((?N o simpnum) ` ?B)" by (simp add:image_comp)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5150
  also have "\<dots> = ?N ` ?B" using simpnum_ci[where bs="real_of_int i #bs"] 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
  5151
  finally have BB': "?N ` set ?B' = ?N ` ?B" .
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5152
  have "?N ` set ?A' = ((?N o simpnum) ` ?A)" by (simp add:image_comp)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5153
  also have "\<dots> = ?N ` ?A" using simpnum_ci[where bs="real_of_int i #bs"] 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
  5154
  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
  5155
  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
  5156
    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
  5157
  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
  5158
    by simp
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  5159
  { 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
  5160
    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
  5161
      using qBd by (auto simp add: Let_def unit_def)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5162
    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
  5163
      and bn: "\<forall>b\<in> set B. numbound0 b" by simp+
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  5164
    with pq_ex dp uq dd lq q d have ?thes by simp }
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5165
  moreover
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  5166
  { 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
  5167
    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
  5168
      using qBd by (auto simp add: Let_def unit_def)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5169
    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
  5170
      and bn: "\<forall>b\<in> set B. numbound0 b" by simp+
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5171
    from mirror_ex[OF lq] pq_ex 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
  5172
    have pqm_eq:"(\<exists> (x::int). ?I x p) = (\<exists> (x::int). ?I x q)" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5173
    from lq uq q mirror_d_\<beta> [where p="?q" and bs="bs" and a="real_of_int i"]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5174
    have lq': "iszlfm q (real_of_int i#bs)" and uq: "d_\<beta> q 1" by auto
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5175
    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
  5176
    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
  5177
  }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5178
  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
  5179
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
  5180
    (* 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
  5181
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  5182
definition cooper :: "fm \<Rightarrow> fm" where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5183
  "cooper p \<equiv>
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  5184
  (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
  5185
       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
  5186
       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
  5187
   in if md = T then T else
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5188
    (let qd = evaldjf (\<lambda> t. simpfm (subst0 t q))
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5189
                               (remdups (map (\<lambda> (b,j). simpnum (Add b (C j)))
24336
fff40259f336 removed allpairs
nipkow
parents: 24249
diff changeset
  5190
                                            [(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
  5191
     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
  5192
lemma cooper: assumes qf: "qfree p"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5193
  shows "((\<exists> (x::int). Ifm (real_of_int x#bs) p) = (Ifm bs (cooper p))) \<and> qfree (cooper 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
  5194
  (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
  5195
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
  5196
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5197
  let ?I = "\<lambda> (x::int) p. Ifm (real_of_int 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
  5198
  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
  5199
  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
  5200
  let ?d = "snd (snd (unit p))"
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  5201
  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
  5202
  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
  5203
  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
  5204
  let ?md = "evaldjf (\<lambda> j. simpfm (subst0 (C j) ?smq)) ?js"
26935
ee6bcb1b8953 avoid undeclared variables within proofs;
wenzelm
parents: 26932
diff changeset
  5205
  fix i
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5206
  let ?N = "\<lambda> t. Inum (real_of_int (i::int)#bs) t"
24336
fff40259f336 removed allpairs
nipkow
parents: 24249
diff changeset
  5207
  let ?bjs = "[(b,j). b\<leftarrow>?B,j\<leftarrow>?js]"
fff40259f336 removed allpairs
nipkow
parents: 24249
diff changeset
  5208
  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
  5209
  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
  5210
  have qbf:"unit p = (?q,?B,?d)" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5211
  from unit[OF qf qbf] have pq_ex: "(\<exists>(x::int). ?I x p) = (\<exists> (x::int). ?I x ?q)" and
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5212
    B:"?N ` set ?B = ?N ` set (\<beta> ?q)" and
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5213
    uq:"d_\<beta> ?q 1" and dd: "d_\<delta> ?q ?d" and dp: "?d > 0" and
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5214
    lq: "iszlfm ?q (real_of_int i#bs)" 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
  5215
    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
  5216
  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
  5217
  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
  5218
  have jsnb: "\<forall> j \<in> set ?js. numbound0 (C j)" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5219
  hence "\<forall> j\<in> set ?js. bound0 (subst0 (C j) ?smq)"
23264
324622260d29 Added 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
    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
  5221
  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
  5222
    by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5223
  from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp
24336
fff40259f336 removed allpairs
nipkow
parents: 24249
diff changeset
  5224
  from Bn jsnb have "\<forall> (b,j) \<in> set ?bjs. numbound0 (Add b (C j))"
fff40259f336 removed allpairs
nipkow
parents: 24249
diff changeset
  5225
    by simp
fff40259f336 removed allpairs
nipkow
parents: 24249
diff changeset
  5226
  hence "\<forall> (b,j) \<in> set ?bjs. numbound0 (simpnum (Add b (C j)))"
fff40259f336 removed allpairs
nipkow
parents: 24249
diff changeset
  5227
    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
  5228
  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
  5229
  hence "\<forall> t \<in> set (remdups ?sbjs). bound0 (subst0 t ?q)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5230
    using subst0_bound0[OF qfq] 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
  5231
  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
  5232
    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
  5233
  from evaldjf_bound0 [OF th'] have qdb: "bound0 ?qd" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5234
  from mdb qdb
23264
324622260d29 Added 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
  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
  5236
  from trans [OF pq_ex cp_thm'[OF lq uq dd dp]] B
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5237
  have "?lhs = (\<exists> j\<in> {1.. ?d}. ?I j ?mq \<or> (\<exists> b\<in> ?N ` set ?B. Ifm ((b+ real_of_int j)#bs) ?q))" by auto
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5238
  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_of_int j)#bs) ?q))" by auto
24336
fff40259f336 removed allpairs
nipkow
parents: 24249
diff changeset
  5239
  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
  5240
  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)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5241
  also  have "\<dots>= ((\<exists> j\<in> set ?js. ?I j ?smq) \<or> (\<exists> t \<in> set ?sbjs. Ifm (?N t #bs) ?q))"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5242
    by (auto simp add: split_def)
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  5243
  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
  5244
    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
  5245
  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
  5246
  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
  5247
  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
  5248
  {assume mdT: "?md = T"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5249
    hence cT:"cooper p = 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
  5250
      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
  5251
    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
  5252
    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
  5253
    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
  5254
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5255
  {assume mdT: "?md \<noteq> T" hence "cooper p = decr (disj ?md ?qd)"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5256
      by (simp only: cooper_def unit_def split_def Let_def if_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
  5257
    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
  5258
  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
  5259
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
  5260
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5261
lemma DJcooper:
23264
324622260d29 Added 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
  assumes qf: "qfree p"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5263
  shows "((\<exists> (x::int). Ifm (real_of_int x#bs) p) = (Ifm bs (DJ cooper p))) \<and> qfree (DJ cooper 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
  5264
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
  5265
  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
  5266
  from DJ_qf[OF cqf] qf have thqf:"qfree (DJ cooper p)" by blast
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5267
  have "Ifm bs (DJ cooper p) = (\<exists> q\<in> set (disjuncts p). Ifm bs (cooper 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
  5268
     by (simp add: DJ_def evaldjf_ex)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5269
  also have "\<dots> = (\<exists> q \<in> set(disjuncts p). \<exists> (x::int). Ifm (real_of_int x#bs)  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
  5270
    using cooper disjuncts_qf[OF qf] by blast
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5271
  also have "\<dots> = (\<exists> (x::int). Ifm (real_of_int x#bs) p)" by (induct p rule: disjuncts.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
  5272
  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
  5273
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
  5274
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5275
    (* 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
  5276
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5277
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
  5278
  shows "Ifm (a#bs) (\<sigma>_\<rho> p (t,c)) = Ifm (a#bs) (\<sigma>_\<rho> p (t',c))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5279
  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
  5280
  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
  5281
324622260d29 Added 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
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
  5283
  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
  5284
  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
  5285
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5286
lemma \<rho>_cong: assumes lp: "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
  5287
  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
  5288
  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
  5289
  (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
  5290
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
  5291
  let ?d = "\<delta> p"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5292
  assume ?lhs then obtain e c j where ecR: "(e,c) \<in> R" and jD:"j \<in> {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
  5293
    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
  5294
  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
  5295
  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
  5296
  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
  5297
  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
  5298
    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
  5299
  from ee' have tt': "Inum (a#bs) (Add e (C j)) = Inum (a#bs) (Add e' (C j))" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5300
23264
324622260d29 Added 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
  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
  5302
  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
  5303
    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
  5304
  (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
  5305
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
  5306
  let ?d = "\<delta> p"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5307
  assume ?rhs then obtain e c j where ecR: "(e,c) \<in> set (\<rho> p)" and jD:"j \<in> {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
  5308
    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
  5309
  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
  5310
  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
  5311
  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
  5312
  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
  5313
    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
  5314
  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
  5315
  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
  5316
  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
  5317
    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
  5318
  (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
  5319
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
  5320
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5321
lemma rl_thm':
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5322
  assumes lp: "iszlfm p (real_of_int (i::int)#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
  5323
  and R: "(\<lambda>(b,k). (Inum (a#bs) b,k)) ` R =  (\<lambda>(b,k). (Inum (a#bs) b,k)) ` set (\<rho> p)"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5324
  shows "(\<exists> (x::int). Ifm (real_of_int x#bs) p) = ((\<exists> j\<in> {1 .. \<delta> p}. Ifm (real_of_int 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)))))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5325
  using rl_thm[OF lp] \<rho>_cong[OF iszlfm_gen[OF lp, rule_format, where y="a"] R] 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
  5326
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  5327
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
  5328
  "chooset p \<equiv> (let q = zlfm p ; d = \<delta> q;
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5329
             B = remdups (map (\<lambda> (t,k). (simpnum t,k)) (\<rho> q)) ;
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5330
             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
  5331
             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
  5332
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5333
lemma chooset: assumes qf: "qfree p"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5334
  shows "\<And> q B d. chooset p = (q,B,d) \<Longrightarrow>
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5335
     ((\<exists> (x::int). Ifm (real_of_int x#bs) p) =
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5336
      (\<exists> (x::int). Ifm (real_of_int x#bs) q)) \<and>
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5337
      ((\<lambda>(t,k). (Inum (real_of_int i#bs) t,k)) ` set B = (\<lambda>(t,k). (Inum (real_of_int i#bs) t,k)) ` set (\<rho> q)) \<and>
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5338
      (\<delta> q = d) \<and> d >0 \<and> iszlfm q (real_of_int (i::int)#bs) \<and> (\<forall> (e,c)\<in> set B. numbound0 e \<and> 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
  5339
proof-
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5340
  fix q B 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
  5341
  assume qBd: "chooset p = (q,B,d)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5342
  let ?thes = "((\<exists> (x::int). Ifm (real_of_int x#bs) p) =
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5343
             (\<exists> (x::int). Ifm (real_of_int x#bs) q)) \<and> ((\<lambda>(t,k). (Inum (real_of_int i#bs) t,k)) ` set B = (\<lambda>(t,k). (Inum (real_of_int i#bs) t,k)) ` set (\<rho> q)) \<and>
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5344
             (\<delta> q = d) \<and> d >0 \<and> iszlfm q (real_of_int (i::int)#bs) \<and> (\<forall> (e,c)\<in> set B. numbound0 e \<and> c>0)"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5345
  let ?I = "\<lambda> (x::int) p. Ifm (real_of_int 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
  5346
  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
  5347
  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
  5348
  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
  5349
  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
  5350
  let ?B'= "remdups (map ?f (\<rho> ?q))"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5351
  let ?A = "set (\<alpha>_\<rho> ?q)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5352
  let ?A'= "remdups (map ?f (\<alpha>_\<rho> ?q))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5353
  from conjunct1[OF zlfm_I[OF qf, where 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
  5354
  have pp': "\<forall> i. ?I i ?q = ?I i p" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5355
  hence pq_ex:"(\<exists> (x::int). ?I x p) = (\<exists> x. ?I x ?q)" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5356
  from iszlfm_gen[OF conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]], rule_format, where y="real_of_int i"]
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5357
  have lq: "iszlfm ?q (real_of_int (i::int)#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
  5358
  from \<delta>[OF lq] have dp:"?d >0" by blast
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5359
  let ?N = "\<lambda> (t,c). (Inum (real_of_int (i::int)#bs) t,c)"
56154
f0a927235162 more complete set of lemmas wrt. image and composition
haftmann
parents: 55584
diff changeset
  5360
  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
  5361
  also have "\<dots> = ?N ` ?B"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5362
    by(simp add: split_def image_comp simpnum_ci[where bs="real_of_int 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
  5363
  finally have BB': "?N ` set ?B' = ?N ` ?B" .
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5364
  have "?N ` set ?A' = ((?N o ?f) ` ?A)" by (simp add: split_def image_comp)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5365
  also have "\<dots> = ?N ` ?A" using simpnum_ci[where bs="real_of_int i #bs"]
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5366
    by(simp add: split_def image_comp simpnum_ci[where bs="real_of_int 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
  5367
  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
  5368
  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
  5369
    by (simp add: split_def)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5370
  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
  5371
    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
  5372
    {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
  5373
    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
  5374
      using qBd by (auto simp add: Let_def chooset_def)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5375
    with BB' B_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
  5376
      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
  5377
  with pq_ex dp lq q d have ?thes by simp}
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5378
  moreover
23264
324622260d29 Added 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
  {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
  5380
    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
  5381
      using qBd by (auto simp add: Let_def chooset_def)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5382
    with AA' mirror_\<alpha>_\<rho>[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\<rho> q)"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5383
      and bn: "\<forall>(e,c)\<in> set B. numbound0 e \<and> c > 0" by auto
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5384
    from mirror_ex[OF lq] pq_ex 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
  5385
    have pqm_eq:"(\<exists> (x::int). ?I x p) = (\<exists> (x::int). ?I x q)" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5386
    from lq q mirror_l [where p="?q" and bs="bs" and a="real_of_int i"]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5387
    have lq': "iszlfm q (real_of_int i#bs)" 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
  5388
    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
  5389
  }
324622260d29 Added 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
  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
  5391
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
  5392
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  5393
definition stage :: "fm \<Rightarrow> int \<Rightarrow> (num \<times> int) \<Rightarrow> fm" where
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  5394
  "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
  5395
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
  5396
  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
  5397
  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
  5398
324622260d29 Added 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
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
  5400
  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
  5401
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
  5402
  let ?f = "\<lambda> j. simpfm (\<sigma> p c (Add e (C j)))"
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  5403
  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
  5404
  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
  5405
    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
  5406
    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
  5407
    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
  5408
    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
  5409
  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
  5410
  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
  5411
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
  5412
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  5413
definition redlove :: "fm \<Rightarrow> fm" where
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5414
  "redlove p \<equiv>
23264
324622260d29 Added 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
  (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
  5416
       mq = simpfm (minusinf q);
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  5417
       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
  5418
   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
  5419
    (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
  5420
     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
  5421
324622260d29 Added 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
lemma redlove: assumes qf: "qfree p"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5423
  shows "((\<exists> (x::int). Ifm (real_of_int x#bs) p) = (Ifm bs (redlove p))) \<and> qfree (redlove 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
  5424
  (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
  5425
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
  5426
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5427
  let ?I = "\<lambda> (x::int) p. Ifm (real_of_int 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
  5428
  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
  5429
  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
  5430
  let ?d = "snd (snd (chooset p))"
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  5431
  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
  5432
  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
  5433
  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
  5434
  let ?md = "evaldjf (\<lambda> j. simpfm (subst0 (C j) ?smq)) ?js"
26935
ee6bcb1b8953 avoid undeclared variables within proofs;
wenzelm
parents: 26932
diff changeset
  5435
  fix i
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5436
  let ?N = "\<lambda> (t,k). (Inum (real_of_int (i::int)#bs) 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
  5437
  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
  5438
  have qbf:"chooset p = (?q,?B,?d)" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5439
  from chooset[OF qf qbf] have pq_ex: "(\<exists>(x::int). ?I x p) = (\<exists> (x::int). ?I x ?q)" and
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5440
    B:"?N ` set ?B = ?N ` set (\<rho> ?q)" and dd: "\<delta> ?q = ?d" and dp: "?d > 0" and
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5441
    lq: "iszlfm ?q (real_of_int i#bs)" 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
  5442
    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
  5443
  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
  5444
  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
  5445
  have jsnb: "\<forall> j \<in> set ?js. numbound0 (C j)" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5446
  hence "\<forall> j\<in> set ?js. bound0 (subst0 (C j) ?smq)"
23264
324622260d29 Added 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
    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
  5448
  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
  5449
    by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5450
  from evaldjf_bound0[OF th] have mdb: "bound0 ?md" 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
  5451
  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
  5452
  from evaldjf_bound0[OF th]  have qdb: "bound0 ?qd" .
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5453
  from mdb qdb
23264
324622260d29 Added 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
  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
  5455
  from trans [OF pq_ex rl_thm'[OF lq B]] dd
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5456
  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_of_int i#bs) (\<sigma> ?q c (Add e (C j)))))" by auto
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5457
  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
  5458
    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
  5459
  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
  5460
    by (simp add: evaldjf_ex subst0_I[OF qfmq])
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5461
  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
  5462
  also have "\<dots> = (?I i (disj ?md ?qd))" by simp
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5463
  also have "\<dots> = (Ifm bs (decr (disj ?md ?qd)))" by (simp only: decr [OF mdqdb])
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5464
  finally have mdqd2: "?lhs = (Ifm bs (decr (disj ?md ?qd)))" .
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5465
  {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
  5466
    hence cT:"redlove p = T" by (simp add: redlove_def Let_def chooset_def split_def)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5467
    from mdT have lhs:"?lhs" using mdqd 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
  5468
    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
  5469
    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
  5470
  moreover
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5471
  {assume mdT: "?md \<noteq> T" hence "redlove p = decr (disj ?md ?qd)"
23264
324622260d29 Added 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
      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
  5473
    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
  5474
  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
  5475
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
  5476
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5477
lemma DJredlove:
23264
324622260d29 Added 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
  assumes qf: "qfree p"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5479
  shows "((\<exists> (x::int). Ifm (real_of_int x#bs) p) = (Ifm bs (DJ redlove p))) \<and> qfree (DJ redlove 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
  5480
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
  5481
  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
  5482
  from DJ_qf[OF cqf] qf have thqf:"qfree (DJ redlove p)" by blast
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5483
  have "Ifm bs (DJ redlove p) = (\<exists> q\<in> set (disjuncts p). Ifm bs (redlove 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
  5484
     by (simp add: DJ_def evaldjf_ex)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5485
  also have "\<dots> = (\<exists> q \<in> set(disjuncts p). \<exists> (x::int). Ifm (real_of_int x#bs)  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
  5486
    using redlove disjuncts_qf[OF qf] by blast
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5487
  also have "\<dots> = (\<exists> (x::int). Ifm (real_of_int x#bs) p)" by (induct p rule: disjuncts.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
  5488
  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
  5489
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
  5490
324622260d29 Added 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
324622260d29 Added 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
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
  5493
  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
  5494
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
  5495
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5496
definition mircfr :: "fm \<Rightarrow> fm" where
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5497
  "mircfr = DJ cooper o ferrack01 o simpfm o exsplit"
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5498
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5499
definition mirlfr :: "fm \<Rightarrow> fm" where
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5500
  "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
  5501
324622260d29 Added 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
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
  5503
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
  5504
  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
  5505
  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
  5506
  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
  5507
  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
  5508
    let ?es = "(And (And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (simpfm (exsplit p)))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5509
    have "?rhs = (\<exists> (i::int). \<exists> x. Ifm (x#real_of_int i#bs) ?es)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5510
      using splitex[OF qf] by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5511
    with ferrack01[OF simpfm_qf[OF exsplit_qf[OF qf]]] have th1: "?rhs = (\<exists> (i::int). Ifm (real_of_int i#bs) (ferrack01 (simpfm (exsplit p))))" and qf':"qfree (ferrack01 (simpfm (exsplit 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
  5512
    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
  5513
  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
  5514
qed
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5515
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5516
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
  5517
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
  5518
  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
  5519
  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
  5520
  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
  5521
  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
  5522
    let ?es = "(And (And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (simpfm (exsplit p)))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5523
    have "?rhs = (\<exists> (i::int). \<exists> x. Ifm (x#real_of_int i#bs) ?es)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5524
      using splitex[OF qf] by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5525
    with ferrack01[OF simpfm_qf[OF exsplit_qf[OF qf]]] have th1: "?rhs = (\<exists> (i::int). Ifm (real_of_int i#bs) (ferrack01 (simpfm (exsplit p))))" and qf':"qfree (ferrack01 (simpfm (exsplit 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
  5526
    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
  5527
  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
  5528
qed
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5529
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5530
definition mircfrqe:: "fm \<Rightarrow> fm" where
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5531
  "mircfrqe p = qelim (prep p) mircfr"
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5532
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5533
definition mirlfrqe:: "fm \<Rightarrow> fm" where
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5534
  "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
  5535
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5536
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
  5537
  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
  5538
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5539
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
  5540
  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
  5541
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  5542
definition
36870
b897bd9ca71b tuned test problems
haftmann
parents: 36778
diff changeset
  5543
  "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
  5544
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  5545
definition
36870
b897bd9ca71b tuned test problems
haftmann
parents: 36778
diff changeset
  5546
  "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
  5547
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  5548
definition
36870
b897bd9ca71b tuned test problems
haftmann
parents: 36778
diff changeset
  5549
  "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
  5550
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  5551
definition
36870
b897bd9ca71b tuned test problems
haftmann
parents: 36778
diff changeset
  5552
  "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
  5553
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  5554
ML_val \<open>@{code mircfrqe} @{code problem1}\<close>
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  5555
ML_val \<open>@{code mirlfrqe} @{code problem1}\<close>
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  5556
ML_val \<open>@{code mircfrqe} @{code problem2}\<close>
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  5557
ML_val \<open>@{code mirlfrqe} @{code problem2}\<close>
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  5558
ML_val \<open>@{code mircfrqe} @{code problem3}\<close>
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  5559
ML_val \<open>@{code mirlfrqe} @{code problem3}\<close>
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  5560
ML_val \<open>@{code mircfrqe} @{code problem4}\<close>
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  5561
ML_val \<open>@{code mirlfrqe} @{code problem4}\<close>
51272
9c8d63b4b6be prefer stateless 'ML_val' for tests;
wenzelm
parents: 51143
diff changeset
  5562
24249
1f60b45c5f97 renamed keyword "to" to "module_name"
haftmann
parents: 23997
diff changeset
  5563
36531
19f6e3b0d9b6 code_reflect: specify module name directly after keyword
haftmann
parents: 36526
diff changeset
  5564
(*code_reflect Mir
36526
353041483b9b use code_reflect
haftmann
parents: 35416
diff changeset
  5565
  functions mircfrqe mirlfrqe
353041483b9b use code_reflect
haftmann
parents: 35416
diff changeset
  5566
  file "mir.ML"*)
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  5567
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  5568
oracle mirfr_oracle = \<open>
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5569
let
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5570
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5571
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
  5572
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
  5573
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
  5574
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
  5575
fun num_of_term vs (t as Free (xn, xT)) = (case AList.lookup (=) vs t
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5576
     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
  5577
      | SOME n => mk_Bound n)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5578
  | num_of_term vs @{term "of_int (0::int)"} = mk_C 0
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5579
  | num_of_term vs @{term "of_int (1::int)"} = mk_C 1
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5580
  | 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
  5581
  | num_of_term vs @{term "1::real"} = mk_C 1
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54230
diff changeset
  5582
  | 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
  5583
  | num_of_term vs (Bound i) = mk_Bound i
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5584
  | num_of_term vs (@{term "uminus :: real \<Rightarrow> real"} $ t') = @{code Neg} (num_of_term vs t')
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
  5585
  | num_of_term vs (@{term "(+) :: real \<Rightarrow> real \<Rightarrow> real"} $ t1 $ t2) =
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5586
      @{code Add} (num_of_term vs t1, num_of_term vs t2)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
  5587
  | num_of_term vs (@{term "(-) :: real \<Rightarrow> real \<Rightarrow> real"} $ t1 $ t2) =
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5588
      @{code Sub} (num_of_term vs t1, num_of_term vs t2)
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68270
diff changeset
  5589
  | num_of_term vs (@{term "(*) :: real \<Rightarrow> real \<Rightarrow> real"} $ t1 $ t2) =
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5590
      (case (num_of_term vs t1)
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5591
       of @{code C} i => @{code Mul} (i, num_of_term vs t2)
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5592
        | _ => error "num_of_term: unsupported Multiplication")
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5593
  | num_of_term vs (@{term "of_int :: int \<Rightarrow> real"} $ (@{term "numeral :: _ \<Rightarrow> int"} $ t')) =
62342
1cf129590be8 consolidated name
haftmann
parents: 61945
diff changeset
  5594
      mk_C (HOLogic.dest_numeral t')
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5595
  | num_of_term vs (@{term "of_int :: int \<Rightarrow> real"} $ (@{term "- numeral :: _ \<Rightarrow> int"} $ t')) =
62342
1cf129590be8 consolidated name
haftmann
parents: 61945
diff changeset
  5596
      mk_C (~ (HOLogic.dest_numeral t'))
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5597
  | num_of_term vs (@{term "of_int :: int \<Rightarrow> real"} $ (@{term "floor :: real \<Rightarrow> int"} $ t')) =
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5598
      @{code Floor} (num_of_term vs t')
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5599
  | num_of_term vs (@{term "of_int :: int \<Rightarrow> real"} $ (@{term "ceiling :: real \<Rightarrow> int"} $ t')) =
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5600
      @{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
  5601
  | num_of_term vs (@{term "numeral :: _ \<Rightarrow> real"} $ t') =
62342
1cf129590be8 consolidated name
haftmann
parents: 61945
diff changeset
  5602
      mk_C (HOLogic.dest_numeral t')
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54230
diff changeset
  5603
  | num_of_term vs (@{term "- numeral :: _ \<Rightarrow> real"} $ t') =
62342
1cf129590be8 consolidated name
haftmann
parents: 61945
diff changeset
  5604
      mk_C (~ (HOLogic.dest_numeral t'))
28264
e1dae766c108 local @{context};
wenzelm
parents: 27567
diff changeset
  5605
  | 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
  5606
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5607
fun fm_of_term vs @{term True} = @{code T}
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5608
  | fm_of_term vs @{term False} = @{code F}
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
  5609
  | fm_of_term vs (@{term "(<) :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) =
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5610
      @{code Lt} (@{code Sub} (num_of_term vs t1, num_of_term vs t2))
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
  5611
  | fm_of_term vs (@{term "(\<le>) :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) =
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5612
      @{code Le} (@{code Sub} (num_of_term vs t1, num_of_term vs t2))
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
  5613
  | fm_of_term vs (@{term "(=) :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) =
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5614
      @{code Eq} (@{code Sub} (num_of_term vs t1, num_of_term vs t2))
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
  5615
  | fm_of_term vs (@{term "(rdvd)"} $ (@{term "of_int :: int \<Rightarrow> real"} $ (@{term "numeral :: _ \<Rightarrow> int"} $ t1)) $ t2) =
62342
1cf129590be8 consolidated name
haftmann
parents: 61945
diff changeset
  5616
      mk_Dvd (HOLogic.dest_numeral t1, num_of_term vs t2)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
  5617
  | fm_of_term vs (@{term "(rdvd)"} $ (@{term "of_int :: int \<Rightarrow> real"} $ (@{term "- numeral :: _ \<Rightarrow> int"} $ t1)) $ t2) =
62342
1cf129590be8 consolidated name
haftmann
parents: 61945
diff changeset
  5618
      mk_Dvd (~ (HOLogic.dest_numeral t1), num_of_term vs t2)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
  5619
  | fm_of_term vs (@{term "(=) :: bool \<Rightarrow> bool \<Rightarrow> bool"} $ t1 $ t2) =
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5620
      @{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
  5621
  | fm_of_term vs (@{term HOL.conj} $ t1 $ t2) =
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5622
      @{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
  5623
  | fm_of_term vs (@{term HOL.disj} $ t1 $ t2) =
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5624
      @{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
  5625
  | fm_of_term vs (@{term HOL.implies} $ t1 $ t2) =
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5626
      @{code Imp} (fm_of_term vs t1, fm_of_term vs t2)
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5627
  | fm_of_term vs (@{term "Not"} $ t') =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5628
      @{code NOT} (fm_of_term vs t')
38558
32ad17fe2b9c tuned quotes
haftmann
parents: 38549
diff changeset
  5629
  | fm_of_term vs (Const (@{const_name Ex}, _) $ Abs (xn, xT, p)) =
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5630
      @{code E} (fm_of_term (map (fn (v, n) => (v, n + 1)) vs) p)
38558
32ad17fe2b9c tuned quotes
haftmann
parents: 38549
diff changeset
  5631
  | fm_of_term vs (Const (@{const_name All}, _) $ Abs (xn, xT, p)) =
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5632
      @{code A} (fm_of_term (map (fn (v, n) => (v, n + 1)) vs) p)
28264
e1dae766c108 local @{context};
wenzelm
parents: 27567
diff changeset
  5633
  | 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
  5634
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5635
fun term_of_num vs (@{code C} i) = @{term "of_int :: int \<Rightarrow> real"} $
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5636
      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
  5637
  | 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
  5638
      let
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5639
        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
  5640
      in fst (the (find_first (fn (_, q) => m = q) vs)) end
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5641
  | term_of_num vs (@{code Neg} (@{code Floor} (@{code Neg} t'))) =
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5642
      @{term "of_int :: int \<Rightarrow> real"} $ (@{term "ceiling :: real \<Rightarrow> int"} $ term_of_num vs t')
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5643
  | term_of_num vs (@{code Neg} t') = @{term "uminus :: real \<Rightarrow> real"} $ term_of_num vs t'
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
  5644
  | term_of_num vs (@{code Add} (t1, t2)) = @{term "(+) :: real \<Rightarrow> real \<Rightarrow> real"} $
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5645
      term_of_num vs t1 $ term_of_num vs t2
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
  5646
  | term_of_num vs (@{code Sub} (t1, t2)) = @{term "(-) :: real \<Rightarrow> real \<Rightarrow> real"} $
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5647
      term_of_num vs t1 $ term_of_num vs t2
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68270
diff changeset
  5648
  | term_of_num vs (@{code Mul} (i, t2)) = @{term "(*) :: real \<Rightarrow> real \<Rightarrow> real"} $
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5649
      term_of_num vs (@{code C} i) $ term_of_num vs t2
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5650
  | term_of_num vs (@{code Floor} t) = @{term "of_int :: int \<Rightarrow> real"} $ (@{term "floor :: real \<Rightarrow> int"} $ term_of_num vs t)
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5651
  | 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
  5652
  | 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
  5653
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5654
fun term_of_fm vs @{code T} = @{term True}
45740
132a3e1c0fe5 more antiquotations;
wenzelm
parents: 44890
diff changeset
  5655
  | term_of_fm vs @{code F} = @{term False}
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5656
  | term_of_fm vs (@{code Lt} t) =
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
  5657
      @{term "(<) :: real \<Rightarrow> real \<Rightarrow> bool"} $ term_of_num vs t $ @{term "0::real"}
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5658
  | term_of_fm vs (@{code Le} t) =
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
  5659
      @{term "(\<le>) :: real \<Rightarrow> real \<Rightarrow> bool"} $ term_of_num vs t $ @{term "0::real"}
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5660
  | term_of_fm vs (@{code Gt} t) =
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
  5661
      @{term "(<) :: real \<Rightarrow> real \<Rightarrow> bool"} $ @{term "0::real"} $ term_of_num vs t
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5662
  | term_of_fm vs (@{code Ge} t) =
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
  5663
      @{term "(\<le>) :: real \<Rightarrow> real \<Rightarrow> bool"} $ @{term "0::real"} $ term_of_num vs t
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5664
  | term_of_fm vs (@{code Eq} t) =
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
  5665
      @{term "(=) :: real \<Rightarrow> real \<Rightarrow> bool"} $ term_of_num vs t $ @{term "0::real"}
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5666
  | term_of_fm vs (@{code NEq} t) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5667
      term_of_fm vs (@{code NOT} (@{code Eq} t))
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5668
  | term_of_fm vs (@{code Dvd} (i, t)) =
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
  5669
      @{term "(rdvd)"} $ term_of_num vs (@{code C} i) $ term_of_num vs t
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5670
  | term_of_fm vs (@{code NDvd} (i, t)) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5671
      term_of_fm vs (@{code NOT} (@{code Dvd} (i, t)))
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5672
  | term_of_fm vs (@{code NOT} t') =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5673
      HOLogic.Not $ term_of_fm vs t'
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5674
  | term_of_fm vs (@{code And} (t1, t2)) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5675
      HOLogic.conj $ term_of_fm vs t1 $ term_of_fm vs t2
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5676
  | term_of_fm vs (@{code Or} (t1, t2)) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5677
      HOLogic.disj $ term_of_fm vs t1 $ term_of_fm vs t2
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5678
  | term_of_fm vs (@{code Imp}  (t1, t2)) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5679
      HOLogic.imp $ term_of_fm vs t1 $ term_of_fm vs t2
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5680
  | term_of_fm vs (@{code Iff} (t1, t2)) =
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
  5681
      @{term "(=) :: bool \<Rightarrow> bool \<Rightarrow> bool"} $ term_of_fm vs t1 $ term_of_fm vs t2;
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5682
28290
4cc2b6046258 simplified oracle interface;
wenzelm
parents: 28264
diff changeset
  5683
in
60325
6fc771cb42eb clarified context;
wenzelm
parents: 59621
diff changeset
  5684
  fn (ctxt, t) =>
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5685
  let
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 44013
diff changeset
  5686
    val fs = Misc_Legacy.term_frees t;
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 32960
diff changeset
  5687
    val vs = map_index swap fs;
60325
6fc771cb42eb clarified context;
wenzelm
parents: 59621
diff changeset
  5688
    (*If quick_and_dirty then run without proof generation as oracle*)
6fc771cb42eb clarified context;
wenzelm
parents: 59621
diff changeset
  5689
    val qe = if Config.get ctxt quick_and_dirty then @{code mircfrqe} else @{code mirlfrqe};
6fc771cb42eb clarified context;
wenzelm
parents: 59621
diff changeset
  5690
    val t' = term_of_fm vs (qe (fm_of_term vs t));
6fc771cb42eb clarified context;
wenzelm
parents: 59621
diff changeset
  5691
  in Thm.cterm_of ctxt (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, t'))) end
69266
7cc2d66a92a6 proper ML expressions, without trailing semicolons;
wenzelm
parents: 69064
diff changeset
  5692
end
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  5693
\<close>
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5694
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5695
lemmas iff_real_of_int = of_int_eq_iff [where 'a = real, symmetric]
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5696
                         of_int_less_iff [where 'a = real, symmetric]
61652
90c65a811257 MIR decision procedure again working
paulson <lp15@cam.ac.uk>
parents: 61649
diff changeset
  5697
                         of_int_le_iff [where 'a = real, symmetric]
90c65a811257 MIR decision procedure again working
paulson <lp15@cam.ac.uk>
parents: 61649
diff changeset
  5698
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 47432
diff changeset
  5699
ML_file "mir_tac.ML"
47432
e1576d13e933 more standard method setup;
wenzelm
parents: 47142
diff changeset
  5700
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  5701
method_setup mir = \<open>
53168
d998de7f0efc tuned signature;
wenzelm
parents: 51369
diff changeset
  5702
  Scan.lift (Args.mode "no_quantify") >>
47432
e1576d13e933 more standard method setup;
wenzelm
parents: 47142
diff changeset
  5703
    (fn q => fn ctxt => SIMPLE_METHOD' (Mir_Tac.mir_tac ctxt (not q)))
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60325
diff changeset
  5704
\<close> "decision procedure for MIR arithmetic"
61652
90c65a811257 MIR decision procedure again working
paulson <lp15@cam.ac.uk>
parents: 61649
diff changeset
  5705
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5706
lemma "\<forall>x::real. (\<lfloor>x\<rfloor> = \<lceil>x\<rceil> \<longleftrightarrow> (x = real_of_int \<lfloor>x\<rfloor>))"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  5707
  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
  5708
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  5709
lemma "\<forall>x::real. real_of_int (2::int)*x - (real_of_int (1::int)) < real_of_int \<lfloor>x\<rfloor> + real_of_int \<lceil>x\<rceil> \<and> real_of_int \<lfloor>x\<rfloor> + real_of_int \<lceil>x\<rceil>  \<le> real_of_int (2::int)*x + (real_of_int (1::int))"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  5710
  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
  5711
58909
f323497583d1 more symbols;
wenzelm
parents: 58410
diff changeset
  5712
lemma "\<forall>x::real. 2*\<lfloor>x\<rfloor> \<le> \<lfloor>2*x\<rfloor> \<and> \<lfloor>2*x\<rfloor> \<le> 2*\<lfloor>x+1\<rfloor>"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
  5713
  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
  5714
58909
f323497583d1 more symbols;
wenzelm
parents: 58410
diff changeset
  5715
lemma "\<forall>x::real. \<exists>y \<le> x. (\<lfloor>x\<rfloor> = \<lceil>y\<rceil>)"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  5716
  by mir
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  5717
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  5718
lemma "\<forall>(x::real) (y::real). \<lfloor>x\<rfloor> = \<lfloor>y\<rfloor> \<longrightarrow> 0 \<le> \<bar>y - x\<bar> \<and> \<bar>y - x\<bar> \<le> 1"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  5719
  by mir
61652
90c65a811257 MIR decision procedure again working
paulson <lp15@cam.ac.uk>
parents: 61649
diff changeset
  5720
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5721
end