src/HOL/Decision_Procs/Ferrack.thy
author haftmann
Thu, 19 Jun 2025 17:15:40 +0200
changeset 82734 89347c0cc6a3
parent 82292 5d91cca0aaf3
permissions -rw-r--r--
treat map_filter similar to list_all, list_ex, list_ex1
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
30439
57c68b3af2ea Updated paths in Decision_Procs comments and NEWS
hoelzl
parents: 30042
diff changeset
     1
(*  Title:      HOL/Decision_Procs/Ferrack.thy
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
     2
    Author:     Amine Chaieb
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
     3
*)
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
     4
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
     5
theory Ferrack
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41842
diff changeset
     6
imports Complex_Main Dense_Linear_Order DP_Library
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
     7
  "HOL-Library.Code_Target_Numeral"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
     8
begin
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
     9
61586
5197a2ecb658 isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
    10
section \<open>Quantifier elimination for \<open>\<real> (0, 1, +, <)\<close>\<close>
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
    11
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
    12
  (*********************************************************************************)
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
    13
  (****                            SHADOW SYNTAX AND SEMANTICS                  ****)
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
    14
  (*********************************************************************************)
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
    15
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    16
datatype (plugins del: size) num = C int | Bound nat | CN nat int num | Neg num | Add num num| Sub num num
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
    17
  | Mul int num
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
    18
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    19
instantiation num :: size
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    20
begin
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    21
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    22
primrec size_num :: "num \<Rightarrow> nat"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
    23
where
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    24
  "size_num (C c) = 1"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    25
| "size_num (Bound n) = 1"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    26
| "size_num (Neg a) = 1 + size_num a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    27
| "size_num (Add a b) = 1 + size_num a + size_num b"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    28
| "size_num (Sub a b) = 3 + size_num a + size_num b"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    29
| "size_num (Mul c a) = 1 + size_num a"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    30
| "size_num (CN n c a) = 3 + size_num a "
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    31
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    32
instance ..
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    33
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    34
end
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
    35
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
    36
  (* Semantics of numeral terms (num) *)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
    37
primrec Inum :: "real list \<Rightarrow> num \<Rightarrow> real"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
    38
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
    39
  "Inum bs (C c) = (real_of_int c)"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
    40
| "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
    41
| "Inum bs (CN n c a) = (real_of_int c) * (bs!n) + (Inum bs a)"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
    42
| "Inum bs (Neg a) = -(Inum bs a)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
    43
| "Inum bs (Add a b) = Inum bs a + Inum bs b"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
    44
| "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
    45
| "Inum bs (Mul c a) = (real_of_int c) * Inum bs a"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
    46
    (* FORMULAE *)
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    47
datatype (plugins del: size) fm  =
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
    48
  T| F| Lt num| Le num| Gt num| Ge num| Eq num| NEq num|
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
    49
  Not fm| And fm fm|  Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
    50
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    51
instantiation fm :: size
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    52
begin
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
    53
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    54
primrec size_fm :: "fm \<Rightarrow> nat"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
    55
where
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
    56
  "size_fm (Not p) = 1 + size_fm p"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    57
| "size_fm (And p q) = 1 + size_fm p + size_fm q"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    58
| "size_fm (Or p q) = 1 + size_fm p + size_fm q"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    59
| "size_fm (Imp p q) = 3 + size_fm p + size_fm q"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    60
| "size_fm (Iff p q) = 3 + 2 * (size_fm p + size_fm q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    61
| "size_fm (E p) = 1 + size_fm p"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    62
| "size_fm (A p) = 4 + size_fm p"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    63
| "size_fm T = 1"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    64
| "size_fm F = 1"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    65
| "size_fm (Lt _) = 1"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    66
| "size_fm (Le _) = 1"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    67
| "size_fm (Gt _) = 1"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    68
| "size_fm (Ge _) = 1"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    69
| "size_fm (Eq _) = 1"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    70
| "size_fm (NEq _) = 1"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
    71
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    72
instance ..
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    73
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    74
end
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    75
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    76
lemma size_fm_pos [simp]: "size p > 0" for p :: fm
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    77
  by (induct p) simp_all
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
    78
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
    79
  (* Semantics of formulae (fm) *)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
    80
primrec Ifm ::"real list \<Rightarrow> fm \<Rightarrow> bool"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
    81
where
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
    82
  "Ifm bs T = True"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
    83
| "Ifm bs F = False"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
    84
| "Ifm bs (Lt a) = (Inum bs a < 0)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
    85
| "Ifm bs (Gt a) = (Inum bs a > 0)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
    86
| "Ifm bs (Le a) = (Inum bs a \<le> 0)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
    87
| "Ifm bs (Ge a) = (Inum bs a \<ge> 0)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
    88
| "Ifm bs (Eq a) = (Inum bs a = 0)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
    89
| "Ifm bs (NEq a) = (Inum bs a \<noteq> 0)"
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
    90
| "Ifm bs (Not p) = (\<not> (Ifm bs p))"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
    91
| "Ifm bs (And p q) = (Ifm bs p \<and> Ifm bs q)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
    92
| "Ifm bs (Or p q) = (Ifm bs p \<or> Ifm bs q)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
    93
| "Ifm bs (Imp p q) = ((Ifm bs p) \<longrightarrow> (Ifm bs q))"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
    94
| "Ifm bs (Iff p q) = (Ifm bs p = Ifm bs q)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
    95
| "Ifm bs (E p) = (\<exists>x. Ifm (x#bs) p)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
    96
| "Ifm bs (A p) = (\<forall>x. Ifm (x#bs) p)"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
    97
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
    98
lemma IfmLeSub: "\<lbrakk> Inum bs s = s' ; Inum bs t = t' \<rbrakk> \<Longrightarrow> Ifm bs (Le (Sub s t)) = (s' \<le> t')"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
    99
  by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   100
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   101
lemma IfmLtSub: "\<lbrakk> Inum bs s = s' ; Inum bs t = t' \<rbrakk> \<Longrightarrow> Ifm bs (Lt (Sub s t)) = (s' < t')"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   102
  by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   103
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   104
lemma IfmEqSub: "\<lbrakk> Inum bs s = s' ; Inum bs t = t' \<rbrakk> \<Longrightarrow> Ifm bs (Eq (Sub s t)) = (s' = t')"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   105
  by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   106
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
   107
lemma IfmNot: " (Ifm bs p = P) \<Longrightarrow> (Ifm bs (Not p) = (\<not>P))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   108
  by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   109
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   110
lemma IfmAnd: " \<lbrakk> Ifm bs p = P ; Ifm bs q = Q\<rbrakk> \<Longrightarrow> (Ifm bs (And p q) = (P \<and> Q))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   111
  by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   112
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   113
lemma IfmOr: " \<lbrakk> Ifm bs p = P ; Ifm bs q = Q\<rbrakk> \<Longrightarrow> (Ifm bs (Or p q) = (P \<or> Q))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   114
  by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   115
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   116
lemma IfmImp: " \<lbrakk> Ifm bs p = P ; Ifm bs q = Q\<rbrakk> \<Longrightarrow> (Ifm bs (Imp p q) = (P \<longrightarrow> Q))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   117
  by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   118
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   119
lemma IfmIff: " \<lbrakk> Ifm bs p = P ; Ifm bs q = Q\<rbrakk> \<Longrightarrow> (Ifm bs (Iff p q) = (P = Q))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   120
  by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   121
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   122
lemma IfmE: " (!! x. Ifm (x#bs) p = P x) \<Longrightarrow> (Ifm bs (E p) = (\<exists>x. P x))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   123
  by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   124
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   125
lemma IfmA: " (!! x. Ifm (x#bs) p = P x) \<Longrightarrow> (Ifm bs (A p) = (\<forall>x. P x))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   126
  by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   127
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   128
fun not:: "fm \<Rightarrow> fm"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   129
where
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
   130
  "not (Not p) = p"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   131
| "not T = F"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   132
| "not F = T"
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
   133
| "not p = Not p"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   134
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
   135
lemma not[simp]: "Ifm bs (not p) = Ifm bs (Not p)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   136
  by (cases p) auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   137
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   138
definition conj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   139
where
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   140
  "conj p q =
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   141
   (if p = F \<or> q = F then F
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   142
    else if p = T then q
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   143
    else if q = T then p
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   144
    else if p = q then p else And p q)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   145
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   146
lemma conj[simp]: "Ifm bs (conj p q) = Ifm bs (And p q)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   147
  by (cases "p = F \<or> q = F", simp_all add: conj_def) (cases p, simp_all)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   148
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   149
definition disj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   150
where
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   151
  "disj p q =
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   152
   (if p = T \<or> q = T then T
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   153
    else if p = F then q
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   154
    else if q = F then p
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   155
    else if p = q then p else Or p q)"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   156
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   157
lemma disj[simp]: "Ifm bs (disj p q) = Ifm bs (Or p q)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   158
  by (cases "p = T \<or> q = T", simp_all add: disj_def) (cases p, simp_all)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   159
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   160
definition imp :: "fm \<Rightarrow> fm \<Rightarrow> fm"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   161
where
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   162
  "imp p q =
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   163
   (if p = F \<or> q = T \<or> p = q then T
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   164
    else if p = T then q
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   165
    else if q = F then not p
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   166
    else Imp p q)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   167
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   168
lemma imp[simp]: "Ifm bs (imp p q) = Ifm bs (Imp p q)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   169
  by (cases "p = F \<or> q = T") (simp_all add: imp_def)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   170
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   171
definition iff :: "fm \<Rightarrow> fm \<Rightarrow> fm"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   172
where
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   173
  "iff p q =
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   174
   (if p = q then T
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
   175
    else if p = Not q \<or> Not p = q then F
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   176
    else if p = F then not q
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   177
    else if q = F then not p
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   178
    else if p = T then q
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   179
    else if q = T then p
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   180
    else Iff p q)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   181
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   182
lemma iff[simp]: "Ifm bs (iff p q) = Ifm bs (Iff p q)"
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
   183
  by (unfold iff_def, cases "p = q", simp, cases "p = Not q", simp) (cases "Not p = q", auto)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   184
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   185
lemma conj_simps:
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   186
  "conj F Q = F"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   187
  "conj P F = F"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   188
  "conj T Q = Q"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   189
  "conj P T = P"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   190
  "conj P P = P"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   191
  "P \<noteq> T \<Longrightarrow> P \<noteq> F \<Longrightarrow> Q \<noteq> T \<Longrightarrow> Q \<noteq> F \<Longrightarrow> P \<noteq> Q \<Longrightarrow> conj P Q = And P Q"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   192
  by (simp_all add: conj_def)
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   193
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   194
lemma disj_simps:
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   195
  "disj T Q = T"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   196
  "disj P T = T"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   197
  "disj F Q = Q"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   198
  "disj P F = P"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   199
  "disj P P = P"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   200
  "P \<noteq> T \<Longrightarrow> P \<noteq> F \<Longrightarrow> Q \<noteq> T \<Longrightarrow> Q \<noteq> F \<Longrightarrow> P \<noteq> Q \<Longrightarrow> disj P Q = Or P Q"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   201
  by (simp_all add: disj_def)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   202
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   203
lemma imp_simps:
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   204
  "imp F Q = T"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   205
  "imp P T = T"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   206
  "imp T Q = Q"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   207
  "imp P F = not P"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   208
  "imp P P = T"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   209
  "P \<noteq> T \<Longrightarrow> P \<noteq> F \<Longrightarrow> P \<noteq> Q \<Longrightarrow> Q \<noteq> T \<Longrightarrow> Q \<noteq> F \<Longrightarrow> imp P Q = Imp P Q"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   210
  by (simp_all add: imp_def)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   211
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
   212
lemma trivNot: "p \<noteq> Not p" "Not p \<noteq> p"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   213
  by (induct p) auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   214
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   215
lemma iff_simps:
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   216
  "iff p p = T"
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
   217
  "iff p (Not p) = F"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
   218
  "iff (Not p) p = F"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   219
  "iff p F = not p"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   220
  "iff F p = not p"
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
   221
  "p \<noteq> Not T \<Longrightarrow> iff T p = p"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
   222
  "p\<noteq> Not T \<Longrightarrow> iff p T = p"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
   223
  "p\<noteq>q \<Longrightarrow> p\<noteq> Not q \<Longrightarrow> q\<noteq> Not p \<Longrightarrow> p\<noteq> F \<Longrightarrow> q\<noteq> F \<Longrightarrow> p \<noteq> T \<Longrightarrow> q \<noteq> T \<Longrightarrow> iff p q = Iff p q"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
   224
  using trivNot
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   225
  by (simp_all add: iff_def, cases p, auto)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   226
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   227
  (* Quantifier freeness *)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   228
fun qfree:: "fm \<Rightarrow> bool"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   229
where
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   230
  "qfree (E p) = False"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   231
| "qfree (A p) = False"
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
   232
| "qfree (Not p) = qfree p"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   233
| "qfree (And p q) = (qfree p \<and> qfree q)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   234
| "qfree (Or  p q) = (qfree p \<and> qfree q)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   235
| "qfree (Imp p q) = (qfree p \<and> qfree q)"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   236
| "qfree (Iff p q) = (qfree p \<and> qfree q)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   237
| "qfree p = True"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   238
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   239
  (* Boundedness and substitution *)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   240
primrec numbound0:: "num \<Rightarrow> bool" (* a num is INDEPENDENT of Bound 0 *)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   241
where
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   242
  "numbound0 (C c) = True"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   243
| "numbound0 (Bound n) = (n > 0)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   244
| "numbound0 (CN n c a) = (n \<noteq> 0 \<and> numbound0 a)"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   245
| "numbound0 (Neg a) = numbound0 a"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   246
| "numbound0 (Add a b) = (numbound0 a \<and> numbound0 b)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   247
| "numbound0 (Sub a b) = (numbound0 a \<and> numbound0 b)"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   248
| "numbound0 (Mul i a) = numbound0 a"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   249
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   250
lemma numbound0_I:
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   251
  assumes nb: "numbound0 a"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   252
  shows "Inum (b#bs) a = Inum (b'#bs) a"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   253
  using nb by (induct a) simp_all
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   254
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   255
primrec bound0:: "fm \<Rightarrow> bool" (* A Formula is independent of Bound 0 *)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   256
where
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   257
  "bound0 T = True"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   258
| "bound0 F = True"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   259
| "bound0 (Lt a) = numbound0 a"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   260
| "bound0 (Le a) = numbound0 a"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   261
| "bound0 (Gt a) = numbound0 a"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   262
| "bound0 (Ge a) = numbound0 a"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   263
| "bound0 (Eq a) = numbound0 a"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   264
| "bound0 (NEq a) = numbound0 a"
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
   265
| "bound0 (Not p) = bound0 p"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   266
| "bound0 (And p q) = (bound0 p \<and> bound0 q)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   267
| "bound0 (Or p q) = (bound0 p \<and> bound0 q)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   268
| "bound0 (Imp p q) = ((bound0 p) \<and> (bound0 q))"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   269
| "bound0 (Iff p q) = (bound0 p \<and> bound0 q)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   270
| "bound0 (E p) = False"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   271
| "bound0 (A p) = False"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   272
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   273
lemma bound0_I:
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   274
  assumes bp: "bound0 p"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   275
  shows "Ifm (b#bs) p = Ifm (b'#bs) p"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   276
  using bp numbound0_I[where b="b" and bs="bs" and b'="b'"]
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   277
  by (induct p) auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   278
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   279
lemma not_qf[simp]: "qfree p \<Longrightarrow> qfree (not p)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   280
  by (cases p) auto
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   281
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   282
lemma not_bn[simp]: "bound0 p \<Longrightarrow> bound0 (not p)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   283
  by (cases p) auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   284
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   285
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   286
lemma conj_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (conj p q)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   287
  using conj_def by auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   288
lemma conj_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (conj p q)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   289
  using conj_def by auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   290
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   291
lemma disj_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (disj p q)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   292
  using disj_def by auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   293
lemma disj_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (disj p q)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   294
  using disj_def by auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   295
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   296
lemma imp_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (imp p q)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   297
  using imp_def by (cases "p=F \<or> q=T",simp_all add: imp_def)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   298
lemma imp_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (imp p q)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   299
  using imp_def by (cases "p=F \<or> q=T \<or> p=q",simp_all add: imp_def)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   300
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   301
lemma iff_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (iff p q)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   302
  unfolding iff_def by (cases "p = q") auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   303
lemma iff_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (iff p q)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   304
  using iff_def unfolding iff_def by (cases "p = q") auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   305
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   306
fun decrnum:: "num \<Rightarrow> num"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   307
where
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   308
  "decrnum (Bound n) = Bound (n - 1)"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   309
| "decrnum (Neg a) = Neg (decrnum a)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   310
| "decrnum (Add a b) = Add (decrnum a) (decrnum b)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   311
| "decrnum (Sub a b) = Sub (decrnum a) (decrnum b)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   312
| "decrnum (Mul c a) = Mul c (decrnum a)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   313
| "decrnum (CN n c a) = CN (n - 1) c (decrnum a)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   314
| "decrnum a = a"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   315
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   316
fun decr :: "fm \<Rightarrow> fm"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   317
where
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   318
  "decr (Lt a) = Lt (decrnum a)"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   319
| "decr (Le a) = Le (decrnum a)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   320
| "decr (Gt a) = Gt (decrnum a)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   321
| "decr (Ge a) = Ge (decrnum a)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   322
| "decr (Eq a) = Eq (decrnum a)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   323
| "decr (NEq a) = NEq (decrnum a)"
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
   324
| "decr (Not p) = Not (decr p)"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   325
| "decr (And p q) = conj (decr p) (decr q)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   326
| "decr (Or p q) = disj (decr p) (decr q)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   327
| "decr (Imp p q) = imp (decr p) (decr q)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   328
| "decr (Iff p q) = iff (decr p) (decr q)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   329
| "decr p = p"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   330
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   331
lemma decrnum:
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   332
  assumes nb: "numbound0 t"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   333
  shows "Inum (x # bs) t = Inum bs (decrnum t)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   334
  using nb by (induct t rule: decrnum.induct) simp_all
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   335
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   336
lemma decr:
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   337
  assumes nb: "bound0 p"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   338
  shows "Ifm (x # bs) p = Ifm bs (decr p)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   339
  using nb by (induct p rule: decr.induct) (simp_all add: decrnum)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   340
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   341
lemma decr_qf: "bound0 p \<Longrightarrow> qfree (decr p)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   342
  by (induct p) simp_all
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   343
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   344
fun isatom :: "fm \<Rightarrow> bool" (* test for atomicity *)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   345
where
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   346
  "isatom T = True"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   347
| "isatom F = True"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   348
| "isatom (Lt a) = True"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   349
| "isatom (Le a) = True"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   350
| "isatom (Gt a) = True"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   351
| "isatom (Ge a) = True"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   352
| "isatom (Eq a) = True"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   353
| "isatom (NEq a) = True"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   354
| "isatom p = False"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   355
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   356
lemma bound0_qf: "bound0 p \<Longrightarrow> qfree p"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   357
  by (induct p) simp_all
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   358
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   359
definition djf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   360
where
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   361
  "djf f p q =
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   362
   (if q = T then T
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   363
    else if q = F then f p
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   364
    else (let fp = f p in case fp of T \<Rightarrow> T | F \<Rightarrow> q | _ \<Rightarrow> Or (f p) q))"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   365
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   366
definition evaldjf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   367
  where "evaldjf f ps = foldr (djf f) ps F"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   368
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   369
lemma djf_Or: "Ifm bs (djf f p q) = Ifm bs (Or (f p) q)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   370
  by (cases "q = T", simp add: djf_def, cases "q = F", simp add: djf_def)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   371
    (cases "f p", simp_all add: Let_def djf_def)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   372
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   373
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   374
lemma djf_simps:
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   375
  "djf f p T = T"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   376
  "djf f p F = f p"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   377
  "q \<noteq> T \<Longrightarrow> q \<noteq> F \<Longrightarrow> djf f p q = (let fp = f p in case fp of T \<Rightarrow> T | F \<Rightarrow> q | _ \<Rightarrow> Or (f p) q)"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   378
  by (simp_all add: djf_def)
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   379
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   380
lemma evaldjf_ex: "Ifm bs (evaldjf f ps) \<longleftrightarrow> (\<exists>p \<in> set ps. Ifm bs (f p))"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   381
  by (induct ps) (simp_all add: evaldjf_def djf_Or)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   382
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   383
lemma evaldjf_bound0:
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   384
  assumes nb: "\<forall>x\<in> set xs. bound0 (f x)"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   385
  shows "bound0 (evaldjf f xs)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   386
  using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   387
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   388
lemma evaldjf_qf:
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   389
  assumes nb: "\<forall>x\<in> set xs. qfree (f x)"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   390
  shows "qfree (evaldjf f xs)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   391
  using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   392
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   393
fun disjuncts :: "fm \<Rightarrow> fm list"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   394
where
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   395
  "disjuncts (Or p q) = disjuncts p @ disjuncts q"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   396
| "disjuncts F = []"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   397
| "disjuncts p = [p]"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   398
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   399
lemma disjuncts: "(\<exists>q\<in> set (disjuncts p). Ifm bs q) = Ifm bs p"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   400
  by (induct p rule: disjuncts.induct) auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   401
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   402
lemma disjuncts_nb: "bound0 p \<Longrightarrow> \<forall>q\<in> set (disjuncts p). bound0 q"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   403
proof -
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   404
  assume nb: "bound0 p"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   405
  then have "list_all bound0 (disjuncts p)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   406
    by (induct p rule: disjuncts.induct) auto
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   407
  then show ?thesis
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   408
    by (simp only: list_all_iff)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   409
qed
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   410
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   411
lemma disjuncts_qf: "qfree p \<Longrightarrow> \<forall>q\<in> set (disjuncts p). qfree q"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   412
proof -
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   413
  assume qf: "qfree p"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   414
  then have "list_all qfree (disjuncts p)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   415
    by (induct p rule: disjuncts.induct) auto
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   416
  then show ?thesis
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   417
    by (simp only: list_all_iff)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   418
qed
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   419
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   420
definition DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   421
  where "DJ f p = evaldjf f (disjuncts p)"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   422
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   423
lemma DJ:
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   424
  assumes fdj: "\<forall>p q. Ifm bs (f (Or p q)) = Ifm bs (Or (f p) (f q))"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   425
    and fF: "f F = F"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   426
  shows "Ifm bs (DJ f p) = Ifm bs (f p)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   427
proof -
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   428
  have "Ifm bs (DJ f p) = (\<exists>q \<in> set (disjuncts p). Ifm bs (f q))"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   429
    by (simp add: DJ_def evaldjf_ex)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   430
  also have "\<dots> = Ifm bs (f p)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   431
    using fdj fF by (induct p rule: disjuncts.induct) auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   432
  finally show ?thesis .
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   433
qed
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   434
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   435
lemma DJ_qf:
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   436
  assumes fqf: "\<forall>p. qfree p \<longrightarrow> qfree (f p)"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   437
  shows "\<forall>p. qfree p \<longrightarrow> qfree (DJ f p) "
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   438
proof clarify
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   439
  fix p
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   440
  assume qf: "qfree p"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   441
  have th: "DJ f p = evaldjf f (disjuncts p)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   442
    by (simp add: DJ_def)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   443
  from disjuncts_qf[OF qf] have "\<forall>q\<in> set (disjuncts p). qfree q" .
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   444
  with fqf have th':"\<forall>q\<in> set (disjuncts p). qfree (f q)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   445
    by blast
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   446
  from evaldjf_qf[OF th'] th show "qfree (DJ f p)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   447
    by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   448
qed
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   449
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   450
lemma DJ_qe:
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   451
  assumes qe: "\<forall>bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm bs (qe p) = Ifm bs (E p))"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   452
  shows "\<forall>bs p. qfree p \<longrightarrow> qfree (DJ qe p) \<and> (Ifm bs ((DJ qe p)) = Ifm bs (E p))"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   453
proof clarify
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   454
  fix p :: fm
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   455
  fix bs
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   456
  assume qf: "qfree p"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   457
  from qe have qth: "\<forall>p. qfree p \<longrightarrow> qfree (qe p)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   458
    by blast
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   459
  from DJ_qf[OF qth] qf have qfth: "qfree (DJ qe p)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   460
    by auto
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   461
  have "Ifm bs (DJ qe p) \<longleftrightarrow> (\<exists>q\<in> set (disjuncts p). Ifm bs (qe q))"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   462
    by (simp add: DJ_def evaldjf_ex)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   463
  also have "\<dots> \<longleftrightarrow> (\<exists>q \<in> set(disjuncts p). Ifm bs (E q))"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   464
    using qe disjuncts_qf[OF qf] by auto
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   465
  also have "\<dots> = Ifm bs (E p)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   466
    by (induct p rule: disjuncts.induct) auto
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   467
  finally show "qfree (DJ qe p) \<and> Ifm bs (DJ qe p) = Ifm bs (E p)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   468
    using qfth by blast
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   469
qed
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   470
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   471
  (* Simplification *)
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   472
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   473
fun maxcoeff:: "num \<Rightarrow> int"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   474
where
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61610
diff changeset
   475
  "maxcoeff (C i) = \<bar>i\<bar>"
1135b8de26c3 more symbols;
wenzelm
parents: 61610
diff changeset
   476
| "maxcoeff (CN n c t) = max \<bar>c\<bar> (maxcoeff t)"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   477
| "maxcoeff t = 1"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   478
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   479
lemma maxcoeff_pos: "maxcoeff t \<ge> 0"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   480
  by (induct t rule: maxcoeff.induct, auto)
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   481
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   482
fun numgcdh:: "num \<Rightarrow> int \<Rightarrow> int"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   483
where
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30684
diff changeset
   484
  "numgcdh (C i) = (\<lambda>g. gcd i g)"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   485
| "numgcdh (CN n c t) = (\<lambda>g. gcd c (numgcdh t g))"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   486
| "numgcdh t = (\<lambda>g. 1)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   487
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   488
definition numgcd :: "num \<Rightarrow> int"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   489
  where "numgcd t = numgcdh t (maxcoeff t)"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   490
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   491
fun reducecoeffh:: "num \<Rightarrow> int \<Rightarrow> num"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   492
where
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   493
  "reducecoeffh (C i) = (\<lambda>g. C (i div g))"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   494
| "reducecoeffh (CN n c t) = (\<lambda>g. CN n (c div g) (reducecoeffh t g))"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   495
| "reducecoeffh t = (\<lambda>g. t)"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   496
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   497
definition reducecoeff :: "num \<Rightarrow> num"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   498
where
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   499
  "reducecoeff t =
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   500
   (let g = numgcd t
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   501
    in if g = 0 then C 0 else if g = 1 then t else reducecoeffh t g)"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   502
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   503
fun dvdnumcoeff:: "num \<Rightarrow> int \<Rightarrow> bool"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   504
where
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   505
  "dvdnumcoeff (C i) = (\<lambda>g. g dvd i)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   506
| "dvdnumcoeff (CN n c t) = (\<lambda>g. g dvd c \<and> dvdnumcoeff t g)"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   507
| "dvdnumcoeff t = (\<lambda>g. False)"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   508
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   509
lemma dvdnumcoeff_trans:
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   510
  assumes gdg: "g dvd g'"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   511
    and dgt':"dvdnumcoeff t g'"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   512
  shows "dvdnumcoeff t g"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   513
  using dgt' gdg
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   514
  by (induct t rule: dvdnumcoeff.induct) (simp_all add: gdg dvd_trans[OF gdg])
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   515
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
   516
declare dvd_trans [trans add]
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   517
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61610
diff changeset
   518
lemma natabs0: "nat \<bar>x\<bar> = 0 \<longleftrightarrow> x = 0"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   519
  by arith
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   520
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   521
lemma numgcd0:
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   522
  assumes g0: "numgcd t = 0"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   523
  shows "Inum bs t = 0"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   524
  using g0[simplified numgcd_def]
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   525
  by (induct t rule: numgcdh.induct) (auto simp add: natabs0 maxcoeff_pos max.absorb2)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   526
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   527
lemma numgcdh_pos:
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   528
  assumes gp: "g \<ge> 0"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   529
  shows "numgcdh t g \<ge> 0"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   530
  using gp by (induct t rule: numgcdh.induct) auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   531
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   532
lemma numgcd_pos: "numgcd t \<ge>0"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   533
  by (simp add: numgcd_def numgcdh_pos maxcoeff_pos)
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   534
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   535
lemma reducecoeffh:
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   536
  assumes gt: "dvdnumcoeff t g"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   537
    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
   538
  shows "real_of_int g *(Inum bs (reducecoeffh t g)) = Inum bs t"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   539
  using gt
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   540
proof (induct t rule: reducecoeffh.induct)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
   541
  case (1 i)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   542
  then have gd: "g dvd i"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   543
    by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   544
  with assms show ?case
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   545
    by (simp add: real_of_int_div[OF gd])
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   546
next
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
   547
  case (2 n c t)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   548
  then have gd: "g dvd c"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   549
    by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   550
  from assms 2 show ?case
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   551
    by (simp add: real_of_int_div[OF gd] algebra_simps)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   552
qed (auto simp add: numgcd_def gp)
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   553
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   554
fun ismaxcoeff:: "num \<Rightarrow> int \<Rightarrow> bool"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   555
where
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61610
diff changeset
   556
  "ismaxcoeff (C i) = (\<lambda>x. \<bar>i\<bar> \<le> x)"
1135b8de26c3 more symbols;
wenzelm
parents: 61610
diff changeset
   557
| "ismaxcoeff (CN n c t) = (\<lambda>x. \<bar>c\<bar> \<le> x \<and> ismaxcoeff t x)"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   558
| "ismaxcoeff t = (\<lambda>x. True)"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   559
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   560
lemma ismaxcoeff_mono: "ismaxcoeff t c \<Longrightarrow> c \<le> c' \<Longrightarrow> ismaxcoeff t c'"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
   561
  by (induct t rule: ismaxcoeff.induct) auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   562
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   563
lemma maxcoeff_ismaxcoeff: "ismaxcoeff t (maxcoeff t)"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   564
proof (induct t rule: maxcoeff.induct)
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   565
  case (2 n c t)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   566
  then have H:"ismaxcoeff t (maxcoeff t)" .
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61610
diff changeset
   567
  have thh: "maxcoeff t \<le> max \<bar>c\<bar> (maxcoeff t)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   568
    by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   569
  from ismaxcoeff_mono[OF H thh] show ?case
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   570
    by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   571
qed simp_all
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   572
67118
ccab07d1196c more simplification rules
haftmann
parents: 66809
diff changeset
   573
lemma zgcd_gt1:
ccab07d1196c more simplification rules
haftmann
parents: 66809
diff changeset
   574
  "\<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
   575
  if "gcd i j > 1" for i j :: int
ccab07d1196c more simplification rules
haftmann
parents: 66809
diff changeset
   576
proof -
ccab07d1196c more simplification rules
haftmann
parents: 66809
diff changeset
   577
  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
   578
    by auto
ccab07d1196c more simplification rules
haftmann
parents: 66809
diff changeset
   579
  with that show ?thesis
ccab07d1196c more simplification rules
haftmann
parents: 66809
diff changeset
   580
    by (auto simp add: not_less)
ccab07d1196c more simplification rules
haftmann
parents: 66809
diff changeset
   581
qed
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   582
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   583
lemma numgcdh0:"numgcdh t m = 0 \<Longrightarrow>  m =0"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   584
  by (induct t rule: numgcdh.induct) auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   585
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   586
lemma dvdnumcoeff_aux:
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   587
  assumes "ismaxcoeff t m"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   588
    and mp: "m \<ge> 0"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   589
    and "numgcdh t m > 1"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   590
  shows "dvdnumcoeff t (numgcdh t m)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   591
  using assms
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   592
proof (induct t rule: numgcdh.induct)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   593
  case (2 n c t)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   594
  let ?g = "numgcdh t m"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   595
  from 2 have th: "gcd c ?g > 1"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   596
    by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   597
  from zgcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"]
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61610
diff changeset
   598
  consider "\<bar>c\<bar> > 1" "?g > 1" | "\<bar>c\<bar> = 0" "?g > 1" | "?g = 0"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   599
    by auto
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   600
  then show ?case
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   601
  proof cases
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   602
    case 1
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   603
    with 2 have th: "dvdnumcoeff t ?g"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   604
      by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   605
    have th': "gcd c ?g dvd ?g"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   606
      by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   607
    from dvdnumcoeff_trans[OF th' th] show ?thesis
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   608
      by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   609
  next
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   610
    case "2'": 2
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   611
    with 2 have th: "dvdnumcoeff t ?g"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   612
      by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   613
    have th': "gcd c ?g dvd ?g"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   614
      by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   615
    from dvdnumcoeff_trans[OF th' th] show ?thesis
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   616
      by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   617
  next
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   618
    case 3
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   619
    then have "m = 0" by (rule numgcdh0)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   620
    with 2 3 show ?thesis by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   621
  qed
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30684
diff changeset
   622
qed auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   623
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   624
lemma dvdnumcoeff_aux2:
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
   625
  assumes "numgcd t > 1"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
   626
  shows "dvdnumcoeff t (numgcd t) \<and> numgcd t > 0"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
   627
  using assms
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   628
proof (simp add: numgcd_def)
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   629
  let ?mc = "maxcoeff t"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   630
  let ?g = "numgcdh t ?mc"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   631
  have th1: "ismaxcoeff t ?mc"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   632
    by (rule maxcoeff_ismaxcoeff)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   633
  have th2: "?mc \<ge> 0"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   634
    by (rule maxcoeff_pos)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   635
  assume H: "numgcdh t ?mc > 1"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   636
  from dvdnumcoeff_aux[OF th1 th2 H] show "dvdnumcoeff t ?g" .
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   637
qed
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   638
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
   639
lemma reducecoeff: "real_of_int (numgcd t) * (Inum bs (reducecoeff t)) = Inum bs t"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   640
proof -
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   641
  let ?g = "numgcd t"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   642
  have "?g \<ge> 0"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   643
    by (simp add: numgcd_pos)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   644
  then consider "?g = 0" | "?g = 1" | "?g > 1" by atomize_elim auto
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   645
  then show ?thesis
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   646
  proof cases
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   647
    case 1
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   648
    then show ?thesis by (simp add: numgcd0)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   649
  next
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   650
    case 2
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   651
    then show ?thesis by (simp add: reducecoeff_def)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   652
  next
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   653
    case g1: 3
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   654
    from dvdnumcoeff_aux2[OF g1] have th1: "dvdnumcoeff t ?g" and g0: "?g > 0"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   655
      by blast+
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   656
    from reducecoeffh[OF th1 g0, where bs="bs"] g1 show ?thesis
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   657
      by (simp add: reducecoeff_def Let_def)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   658
  qed
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   659
qed
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   660
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   661
lemma reducecoeffh_numbound0: "numbound0 t \<Longrightarrow> numbound0 (reducecoeffh t g)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   662
  by (induct t rule: reducecoeffh.induct) auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   663
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   664
lemma reducecoeff_numbound0: "numbound0 t \<Longrightarrow> numbound0 (reducecoeff t)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   665
  using reducecoeffh_numbound0 by (simp add: reducecoeff_def Let_def)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   666
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   667
fun numadd:: "num \<Rightarrow> num \<Rightarrow> num"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   668
where
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   669
  "numadd (CN n1 c1 r1) (CN n2 c2 r2) =
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   670
   (if n1 = n2 then
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   671
    (let c = c1 + c2
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   672
     in (if c = 0 then numadd r1 r2 else CN n1 c (numadd r1 r2)))
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   673
    else if n1 \<le> n2 then (CN n1 c1 (numadd r1 (CN n2 c2 r2)))
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   674
    else (CN n2 c2 (numadd (CN n1 c1 r1) r2)))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   675
| "numadd (CN n1 c1 r1) t = CN n1 c1 (numadd r1 t)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   676
| "numadd t (CN n2 c2 r2) = CN n2 c2 (numadd t r2)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   677
| "numadd (C b1) (C b2) = C (b1 + b2)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   678
| "numadd a b = Add a b"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   679
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   680
lemma numadd [simp]: "Inum bs (numadd t s) = Inum bs (Add t s)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   681
  by (induct t s rule: numadd.induct) (simp_all add: Let_def algebra_simps add_eq_0_iff)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   682
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   683
lemma numadd_nb [simp]: "numbound0 t \<Longrightarrow> numbound0 s \<Longrightarrow> numbound0 (numadd t s)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   684
  by (induct t s rule: numadd.induct) (simp_all add: Let_def)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   685
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   686
fun nummul:: "num \<Rightarrow> int \<Rightarrow> num"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   687
where
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   688
  "nummul (C j) = (\<lambda>i. C (i * j))"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   689
| "nummul (CN n c a) = (\<lambda>i. CN n (i * c) (nummul a i))"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   690
| "nummul t = (\<lambda>i. Mul i t)"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   691
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   692
lemma nummul[simp]: "\<And>i. Inum bs (nummul t i) = Inum bs (Mul i t)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   693
  by (induct t rule: nummul.induct) (auto simp add: algebra_simps)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   694
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   695
lemma nummul_nb[simp]: "\<And>i. numbound0 t \<Longrightarrow> numbound0 (nummul t i)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   696
  by (induct t rule: nummul.induct) auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   697
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   698
definition numneg :: "num \<Rightarrow> num"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   699
  where "numneg t = nummul t (- 1)"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   700
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   701
definition numsub :: "num \<Rightarrow> num \<Rightarrow> num"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   702
  where "numsub s t = (if s = t then C 0 else numadd s (numneg t))"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   703
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   704
lemma numneg[simp]: "Inum bs (numneg t) = Inum bs (Neg t)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   705
  using numneg_def by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   706
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   707
lemma numneg_nb[simp]: "numbound0 t \<Longrightarrow> numbound0 (numneg t)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   708
  using numneg_def by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   709
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   710
lemma numsub[simp]: "Inum bs (numsub a b) = Inum bs (Sub a b)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   711
  using numsub_def by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   712
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   713
lemma numsub_nb[simp]: "\<lbrakk> numbound0 t ; numbound0 s\<rbrakk> \<Longrightarrow> numbound0 (numsub t s)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   714
  using numsub_def by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   715
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   716
primrec simpnum:: "num \<Rightarrow> num"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   717
where
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   718
  "simpnum (C j) = C j"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   719
| "simpnum (Bound n) = CN n 1 (C 0)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   720
| "simpnum (Neg t) = numneg (simpnum t)"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   721
| "simpnum (Add t s) = numadd (simpnum t) (simpnum s)"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   722
| "simpnum (Sub t s) = numsub (simpnum t) (simpnum s)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   723
| "simpnum (Mul i t) = (if i = 0 then C 0 else nummul (simpnum t) i)"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   724
| "simpnum (CN n c t) = (if c = 0 then simpnum t else numadd (CN n c (C 0)) (simpnum t))"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   725
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   726
lemma simpnum_ci[simp]: "Inum bs (simpnum t) = Inum bs t"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   727
  by (induct t) simp_all
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   728
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   729
lemma simpnum_numbound0[simp]: "numbound0 t \<Longrightarrow> numbound0 (simpnum t)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   730
  by (induct t) simp_all
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   731
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   732
fun nozerocoeff:: "num \<Rightarrow> bool"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   733
where
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   734
  "nozerocoeff (C c) = True"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   735
| "nozerocoeff (CN n c t) = (c \<noteq> 0 \<and> nozerocoeff t)"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   736
| "nozerocoeff t = True"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   737
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   738
lemma numadd_nz : "nozerocoeff a \<Longrightarrow> nozerocoeff b \<Longrightarrow> nozerocoeff (numadd a b)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   739
  by (induct a b rule: numadd.induct) (simp_all add: Let_def)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   740
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   741
lemma nummul_nz : "\<And>i. i\<noteq>0 \<Longrightarrow> nozerocoeff a \<Longrightarrow> nozerocoeff (nummul a i)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   742
  by (induct a rule: nummul.induct) (auto simp add: Let_def numadd_nz)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   743
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   744
lemma numneg_nz : "nozerocoeff a \<Longrightarrow> nozerocoeff (numneg a)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   745
  by (simp add: numneg_def nummul_nz)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   746
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   747
lemma numsub_nz: "nozerocoeff a \<Longrightarrow> nozerocoeff b \<Longrightarrow> nozerocoeff (numsub a b)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   748
  by (simp add: numsub_def numneg_nz numadd_nz)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   749
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   750
lemma simpnum_nz: "nozerocoeff (simpnum t)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   751
  by (induct t) (simp_all add: numadd_nz numneg_nz numsub_nz nummul_nz)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   752
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   753
lemma maxcoeff_nz: "nozerocoeff t \<Longrightarrow> maxcoeff t = 0 \<Longrightarrow> t = C 0"
73869
7181130f5872 more default simp rules
haftmann
parents: 69605
diff changeset
   754
  by (induction t rule: maxcoeff.induct) auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   755
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   756
lemma numgcd_nz:
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   757
  assumes nz: "nozerocoeff t"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   758
    and g0: "numgcd t = 0"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   759
  shows "t = C 0"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   760
proof -
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   761
  from g0 have th:"numgcdh t (maxcoeff t) = 0"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   762
    by (simp add: numgcd_def)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   763
  from numgcdh0[OF th] have th:"maxcoeff t = 0" .
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   764
  from maxcoeff_nz[OF nz th] show ?thesis .
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   765
qed
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   766
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   767
definition simp_num_pair :: "(num \<times> int) \<Rightarrow> num \<times> int"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   768
where
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   769
  "simp_num_pair =
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   770
    (\<lambda>(t,n).
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   771
     (if n = 0 then (C 0, 0)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   772
      else
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   773
       (let t' = simpnum t ; g = numgcd t' in
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   774
         if g > 1 then
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   775
          (let g' = gcd n g
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   776
           in if g' = 1 then (t', n) else (reducecoeffh t' g', n div g'))
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   777
         else (t', n))))"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   778
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   779
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
   780
  shows "((\<lambda>(t,n). Inum bs t / real_of_int n) (simp_num_pair (t,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
   781
    ((\<lambda>(t,n). Inum bs t / real_of_int n) (t, n))"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   782
  (is "?lhs = ?rhs")
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   783
proof -
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   784
  let ?t' = "simpnum t"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   785
  let ?g = "numgcd ?t'"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30684
diff changeset
   786
  let ?g' = "gcd n ?g"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   787
  show ?thesis
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   788
  proof (cases "n = 0")
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   789
    case True
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   790
    then show ?thesis
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   791
      by (simp add: Let_def simp_num_pair_def)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   792
  next
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   793
    case nnz: False
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   794
    show ?thesis
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   795
    proof (cases "?g > 1")
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   796
      case False
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   797
      then show ?thesis by (simp add: Let_def simp_num_pair_def)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   798
    next
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   799
      case g1: True
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   800
      then have g0: "?g > 0"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   801
        by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   802
      from g1 nnz have gp0: "?g' \<noteq> 0"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   803
        by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   804
      then have g'p: "?g' > 0"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   805
        using gcd_ge_0_int[where x="n" and y="numgcd ?t'"] by arith
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   806
      then consider "?g' = 1" | "?g' > 1" by arith
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   807
      then show ?thesis
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   808
      proof cases
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   809
        case 1
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   810
        then show ?thesis
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   811
          by (simp add: Let_def simp_num_pair_def)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   812
      next
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   813
        case g'1: 2
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   814
        from dvdnumcoeff_aux2[OF g1] have th1: "dvdnumcoeff ?t' ?g" ..
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32642
diff changeset
   815
        let ?tt = "reducecoeffh ?t' ?g'"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32642
diff changeset
   816
        let ?t = "Inum bs ?tt"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32642
diff changeset
   817
        have gpdg: "?g' dvd ?g" by simp
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   818
        have gpdd: "?g' dvd n" by simp
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32642
diff changeset
   819
        have gpdgp: "?g' dvd ?g'" by simp
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   820
        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
   821
        have th2:"real_of_int ?g' * ?t = Inum bs ?t'"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   822
          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
   823
        from g1 g'1 have "?lhs = ?t / real_of_int (n div ?g')"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   824
          by (simp add: simp_num_pair_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
   825
        also have "\<dots> = (real_of_int ?g' * ?t) / (real_of_int ?g' * (real_of_int (n div ?g')))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   826
          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
   827
        also have "\<dots> = (Inum bs ?t' / real_of_int n)"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 45740
diff changeset
   828
          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
   829
        finally have "?lhs = Inum bs t / real_of_int n"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   830
          by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   831
        then show ?thesis
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   832
          by (simp add: simp_num_pair_def)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   833
      qed
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   834
    qed
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   835
  qed
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   836
qed
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   837
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   838
lemma simp_num_pair_l:
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   839
  assumes tnb: "numbound0 t"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   840
    and np: "n > 0"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   841
    and tn: "simp_num_pair (t, n) = (t', n')"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   842
  shows "numbound0 t' \<and> n' > 0"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   843
proof -
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
   844
  let ?t' = "simpnum t"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   845
  let ?g = "numgcd ?t'"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30684
diff changeset
   846
  let ?g' = "gcd n ?g"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   847
  show ?thesis
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   848
  proof (cases "n = 0")
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   849
    case True
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   850
    then show ?thesis
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   851
      using assms by (simp add: Let_def simp_num_pair_def)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   852
  next
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   853
    case nnz: False
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   854
    show ?thesis
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   855
    proof (cases "?g > 1")
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   856
      case False
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   857
      then show ?thesis
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   858
        using assms by (auto simp add: Let_def simp_num_pair_def)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   859
    next
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   860
      case g1: True
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   861
      then have g0: "?g > 0" by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30684
diff changeset
   862
      from g1 nnz have gp0: "?g' \<noteq> 0" by simp
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   863
      then have g'p: "?g' > 0" using gcd_ge_0_int[where x="n" and y="numgcd ?t'"]
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   864
        by arith
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   865
      then consider "?g'= 1" | "?g' > 1" by arith
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   866
      then show ?thesis
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   867
      proof cases
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   868
        case 1
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   869
        then show ?thesis
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   870
          using assms g1 by (auto simp add: Let_def simp_num_pair_def)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   871
      next
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   872
        case g'1: 2
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32642
diff changeset
   873
        have gpdg: "?g' dvd ?g" by simp
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
   874
        have gpdd: "?g' dvd n" by simp
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32642
diff changeset
   875
        have gpdgp: "?g' dvd ?g'" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32642
diff changeset
   876
        from zdvd_imp_le[OF gpdd np] have g'n: "?g' \<le> n" .
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   877
        from zdiv_mono1[OF g'n g'p, simplified div_self[OF gp0]] have "n div ?g' > 0"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   878
          by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   879
        then show ?thesis
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   880
          using assms g1 g'1
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   881
          by (auto simp add: simp_num_pair_def Let_def reducecoeffh_numbound0)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   882
      qed
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   883
    qed
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   884
  qed
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   885
qed
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   886
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   887
fun simpfm :: "fm \<Rightarrow> fm"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   888
where
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   889
  "simpfm (And p q) = conj (simpfm p) (simpfm q)"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   890
| "simpfm (Or p q) = disj (simpfm p) (simpfm q)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   891
| "simpfm (Imp p q) = imp (simpfm p) (simpfm q)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   892
| "simpfm (Iff p q) = iff (simpfm p) (simpfm q)"
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
   893
| "simpfm (Not p) = not (simpfm p)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   894
| "simpfm (Lt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v < 0) then T else F | _ \<Rightarrow> Lt a')"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   895
| "simpfm (Le a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v \<le> 0)  then T else F | _ \<Rightarrow> Le a')"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   896
| "simpfm (Gt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v > 0)  then T else F | _ \<Rightarrow> Gt a')"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   897
| "simpfm (Ge a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v \<ge> 0)  then T else F | _ \<Rightarrow> Ge a')"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   898
| "simpfm (Eq a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v = 0)  then T else F | _ \<Rightarrow> Eq a')"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   899
| "simpfm (NEq a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v \<noteq> 0)  then T else F | _ \<Rightarrow> NEq a')"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
   900
| "simpfm p = p"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   901
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   902
lemma simpfm: "Ifm bs (simpfm p) = Ifm bs p"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   903
proof (induct p rule: simpfm.induct)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   904
  case (6 a)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   905
  let ?sa = "simpnum a"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   906
  from simpnum_ci have sa: "Inum bs ?sa = Inum bs a"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   907
    by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   908
  consider v where "?sa = C v" | "\<not> (\<exists>v. ?sa = C v)" by blast
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   909
  then show ?case
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   910
  proof cases
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   911
    case 1
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   912
    then show ?thesis using sa by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   913
  next
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   914
    case 2
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   915
    then show ?thesis using sa by (cases ?sa) (simp_all add: Let_def)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   916
  qed
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   917
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   918
  case (7 a)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   919
  let ?sa = "simpnum a"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   920
  from simpnum_ci have sa: "Inum bs ?sa = Inum bs a"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   921
    by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   922
  consider v where "?sa = C v" | "\<not> (\<exists>v. ?sa = C v)" by blast
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   923
  then show ?case
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   924
  proof cases
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   925
    case 1
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   926
    then show ?thesis using sa by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   927
  next
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   928
    case 2
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   929
    then show ?thesis using sa by (cases ?sa) (simp_all add: Let_def)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   930
  qed
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   931
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   932
  case (8 a)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   933
  let ?sa = "simpnum a"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   934
  from simpnum_ci have sa: "Inum bs ?sa = Inum bs a"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   935
    by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   936
  consider v where "?sa = C v" | "\<not> (\<exists>v. ?sa = C v)" by blast
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   937
  then show ?case
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   938
  proof cases
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   939
    case 1
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   940
    then show ?thesis using sa by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   941
  next
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   942
    case 2
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   943
    then show ?thesis using sa by (cases ?sa) (simp_all add: Let_def)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   944
  qed
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   945
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   946
  case (9 a)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   947
  let ?sa = "simpnum a"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   948
  from simpnum_ci have sa: "Inum bs ?sa = Inum bs a"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   949
    by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   950
  consider v where "?sa = C v" | "\<not> (\<exists>v. ?sa = C v)" by blast
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   951
  then show ?case
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   952
  proof cases
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   953
    case 1
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   954
    then show ?thesis using sa by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   955
  next
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   956
    case 2
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   957
    then show ?thesis using sa by (cases ?sa) (simp_all add: Let_def)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   958
  qed
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   959
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   960
  case (10 a)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   961
  let ?sa = "simpnum a"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   962
  from simpnum_ci have sa: "Inum bs ?sa = Inum bs a"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   963
    by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   964
  consider v where "?sa = C v" | "\<not> (\<exists>v. ?sa = C v)" by blast
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   965
  then show ?case
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   966
  proof cases
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   967
    case 1
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   968
    then show ?thesis using sa by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   969
  next
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   970
    case 2
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   971
    then show ?thesis using sa by (cases ?sa) (simp_all add: Let_def)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   972
  qed
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   973
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   974
  case (11 a)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   975
  let ?sa = "simpnum a"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   976
  from simpnum_ci have sa: "Inum bs ?sa = Inum bs a"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   977
    by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   978
  consider v where "?sa = C v" | "\<not> (\<exists>v. ?sa = C v)" by blast
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   979
  then show ?case
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   980
  proof cases
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   981
    case 1
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   982
    then show ?thesis using sa by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   983
  next
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   984
    case 2
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   985
    then show ?thesis using sa by (cases ?sa) (simp_all add: Let_def)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   986
  qed
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   987
qed (induct p rule: simpfm.induct, simp_all)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   988
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   989
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   990
lemma simpfm_bound0: "bound0 p \<Longrightarrow> bound0 (simpfm p)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   991
proof (induct p rule: simpfm.induct)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   992
  case (6 a)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   993
  then have nb: "numbound0 a" by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   994
  then have "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   995
  then show ?case by (cases "simpnum a") (auto simp add: Let_def)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
   996
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   997
  case (7 a)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   998
  then have nb: "numbound0 a" by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
   999
  then have "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1000
  then show ?case by (cases "simpnum a") (auto simp add: Let_def)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1001
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1002
  case (8 a)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1003
  then have nb: "numbound0 a" by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1004
  then have "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1005
  then show ?case by (cases "simpnum a") (auto simp add: Let_def)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1006
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1007
  case (9 a)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1008
  then have nb: "numbound0 a" by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1009
  then have "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1010
  then show ?case by (cases "simpnum a") (auto simp add: Let_def)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1011
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1012
  case (10 a)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1013
  then have nb: "numbound0 a" by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1014
  then have "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1015
  then show ?case by (cases "simpnum a") (auto simp add: Let_def)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1016
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1017
  case (11 a)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1018
  then have nb: "numbound0 a" by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1019
  then have "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1020
  then show ?case by (cases "simpnum a") (auto simp add: Let_def)
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1021
qed (auto simp add: disj_def imp_def iff_def conj_def)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1022
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1023
lemma simpfm_qf: "qfree p \<Longrightarrow> qfree (simpfm p)"
82292
5d91cca0aaf3 Tidied up a few messy proofs
paulson <lp15@cam.ac.uk>
parents: 74397
diff changeset
  1024
  by (induct p rule: simpfm.induct) (auto simp: Let_def split: num.splits)
5d91cca0aaf3 Tidied up a few messy proofs
paulson <lp15@cam.ac.uk>
parents: 74397
diff changeset
  1025
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1026
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1027
fun prep :: "fm \<Rightarrow> fm"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1028
where
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1029
  "prep (E T) = T"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1030
| "prep (E F) = F"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1031
| "prep (E (Or p q)) = disj (prep (E p)) (prep (E q))"
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1032
| "prep (E (Imp p q)) = disj (prep (E (Not p))) (prep (E q))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1033
| "prep (E (Iff p q)) = disj (prep (E (And p q))) (prep (E (And (Not p) (Not q))))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1034
| "prep (E (Not (And p q))) = disj (prep (E (Not p))) (prep (E(Not q)))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1035
| "prep (E (Not (Imp p q))) = prep (E (And p (Not q)))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1036
| "prep (E (Not (Iff p q))) = disj (prep (E (And p (Not q)))) (prep (E(And (Not p) q)))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1037
| "prep (E p) = E (prep p)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1038
| "prep (A (And p q)) = conj (prep (A p)) (prep (A q))"
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1039
| "prep (A p) = prep (Not (E (Not p)))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1040
| "prep (Not (Not p)) = prep p"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1041
| "prep (Not (And p q)) = disj (prep (Not p)) (prep (Not q))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1042
| "prep (Not (A p)) = prep (E (Not p))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1043
| "prep (Not (Or p q)) = conj (prep (Not p)) (prep (Not q))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1044
| "prep (Not (Imp p q)) = conj (prep p) (prep (Not q))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1045
| "prep (Not (Iff p q)) = disj (prep (And p (Not q))) (prep (And (Not p) q))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1046
| "prep (Not p) = not (prep p)"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1047
| "prep (Or p q) = disj (prep p) (prep q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1048
| "prep (And p q) = conj (prep p) (prep q)"
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1049
| "prep (Imp p q) = prep (Or (Not p) q)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1050
| "prep (Iff p q) = disj (prep (And p q)) (prep (And (Not p) (Not q)))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1051
| "prep p = p"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1052
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1053
lemma prep: "\<And>bs. Ifm bs (prep p) = Ifm bs p"
44779
98d597c4193d tuned proofs;
wenzelm
parents: 44013
diff changeset
  1054
  by (induct p rule: prep.induct) auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1055
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1056
  (* Generic quantifier elimination *)
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1057
fun qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1058
where
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1059
  "qelim (E p) = (\<lambda>qe. DJ qe (qelim p qe))"
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1060
| "qelim (A p) = (\<lambda>qe. not (qe ((qelim (Not p) qe))))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1061
| "qelim (Not p) = (\<lambda>qe. not (qelim p qe))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1062
| "qelim (And p q) = (\<lambda>qe. conj (qelim p qe) (qelim q qe))"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1063
| "qelim (Or  p q) = (\<lambda>qe. disj (qelim p qe) (qelim q qe))"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1064
| "qelim (Imp p q) = (\<lambda>qe. imp (qelim p qe) (qelim q qe))"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1065
| "qelim (Iff p q) = (\<lambda>qe. iff (qelim p qe) (qelim q qe))"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1066
| "qelim p = (\<lambda>y. simpfm p)"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1067
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1068
lemma qelim_ci:
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1069
  assumes qe_inv: "\<forall>bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm bs (qe p) = Ifm bs (E p))"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1070
  shows "\<And>bs. qfree (qelim p qe) \<and> (Ifm bs (qelim p qe) = Ifm bs p)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1071
  using qe_inv DJ_qe[OF qe_inv]
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1072
  by (induct p rule: qelim.induct)
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1073
    (auto simp add: simpfm simpfm_qf simp del: simpfm.simps)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1074
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1075
fun minusinf:: "fm \<Rightarrow> fm" (* Virtual substitution of -\<infinity>*)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1076
where
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1077
  "minusinf (And p q) = conj (minusinf p) (minusinf q)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1078
| "minusinf (Or p q) = disj (minusinf p) (minusinf q)"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1079
| "minusinf (Eq  (CN 0 c e)) = F"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1080
| "minusinf (NEq (CN 0 c e)) = T"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1081
| "minusinf (Lt  (CN 0 c e)) = T"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1082
| "minusinf (Le  (CN 0 c e)) = T"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1083
| "minusinf (Gt  (CN 0 c e)) = F"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1084
| "minusinf (Ge  (CN 0 c e)) = F"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1085
| "minusinf p = p"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1086
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1087
fun plusinf:: "fm \<Rightarrow> fm" (* Virtual substitution of +\<infinity>*)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1088
where
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1089
  "plusinf (And p q) = conj (plusinf p) (plusinf q)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1090
| "plusinf (Or p q) = disj (plusinf p) (plusinf q)"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1091
| "plusinf (Eq  (CN 0 c e)) = F"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1092
| "plusinf (NEq (CN 0 c e)) = T"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1093
| "plusinf (Lt  (CN 0 c e)) = F"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1094
| "plusinf (Le  (CN 0 c e)) = F"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1095
| "plusinf (Gt  (CN 0 c e)) = T"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1096
| "plusinf (Ge  (CN 0 c e)) = T"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1097
| "plusinf p = p"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1098
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1099
fun isrlfm :: "fm \<Rightarrow> bool"   (* Linearity test for fm *)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1100
where
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1101
  "isrlfm (And p q) = (isrlfm p \<and> isrlfm q)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1102
| "isrlfm (Or p q) = (isrlfm p \<and> isrlfm q)"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1103
| "isrlfm (Eq  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1104
| "isrlfm (NEq (CN 0 c e)) = (c>0 \<and> numbound0 e)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1105
| "isrlfm (Lt  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1106
| "isrlfm (Le  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1107
| "isrlfm (Gt  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1108
| "isrlfm (Ge  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1109
| "isrlfm p = (isatom p \<and> (bound0 p))"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1110
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1111
  (* splits the bounded from the unbounded part*)
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1112
fun rsplit0 :: "num \<Rightarrow> int \<times> num"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1113
where
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1114
  "rsplit0 (Bound 0) = (1,C 0)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1115
| "rsplit0 (Add a b) = (let (ca,ta) = rsplit0 a; (cb,tb) = rsplit0 b in (ca + cb, Add ta tb))"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1116
| "rsplit0 (Sub a b) = rsplit0 (Add a (Neg b))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1117
| "rsplit0 (Neg a) = (let (c,t) = rsplit0 a in (- c, Neg t))"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1118
| "rsplit0 (Mul c a) = (let (ca,ta) = rsplit0 a in (c * ca, Mul c ta))"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1119
| "rsplit0 (CN 0 c a) = (let (ca,ta) = rsplit0 a in (c + ca, ta))"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1120
| "rsplit0 (CN n c a) = (let (ca,ta) = rsplit0 a in (ca, CN n c ta))"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1121
| "rsplit0 t = (0,t)"
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1122
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60767
diff changeset
  1123
lemma rsplit0: "Inum bs ((case_prod (CN 0)) (rsplit0 t)) = Inum bs t \<and> numbound0 (snd (rsplit0 t))"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1124
proof (induct t rule: rsplit0.induct)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1125
  case (2 a b)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1126
  let ?sa = "rsplit0 a"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1127
  let ?sb = "rsplit0 b"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1128
  let ?ca = "fst ?sa"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1129
  let ?cb = "fst ?sb"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1130
  let ?ta = "snd ?sa"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1131
  let ?tb = "snd ?sb"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1132
  from 2 have nb: "numbound0 (snd(rsplit0 (Add a b)))"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  1133
    by (cases "rsplit0 a") (auto simp add: Let_def split_def)
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60767
diff changeset
  1134
  have "Inum bs ((case_prod (CN 0)) (rsplit0 (Add a b))) =
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60767
diff changeset
  1135
    Inum bs ((case_prod (CN 0)) ?sa)+Inum bs ((case_prod (CN 0)) ?sb)"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1136
    by (simp add: Let_def split_def algebra_simps)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1137
  also have "\<dots> = Inum bs a + Inum bs b"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1138
    using 2 by (cases "rsplit0 a") auto
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1139
  finally show ?case
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1140
    using nb by simp
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 49070
diff changeset
  1141
qed (auto simp add: Let_def split_def algebra_simps, simp add: distrib_left[symmetric])
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1142
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1143
    (* Linearize a formula*)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1144
definition lt :: "int \<Rightarrow> num \<Rightarrow> fm"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1145
where
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1146
  "lt c t = (if c = 0 then (Lt t) else if c > 0 then (Lt (CN 0 c t))
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1147
    else (Gt (CN 0 (-c) (Neg t))))"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1148
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1149
definition le :: "int \<Rightarrow> num \<Rightarrow> fm"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1150
where
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1151
  "le c t = (if c = 0 then (Le t) else if c > 0 then (Le (CN 0 c t))
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1152
    else (Ge (CN 0 (-c) (Neg t))))"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1153
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1154
definition gt :: "int \<Rightarrow> num \<Rightarrow> fm"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1155
where
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1156
  "gt c t = (if c = 0 then (Gt t) else if c > 0 then (Gt (CN 0 c t))
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1157
    else (Lt (CN 0 (-c) (Neg t))))"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1158
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1159
definition ge :: "int \<Rightarrow> num \<Rightarrow> fm"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1160
where
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1161
  "ge c t = (if c = 0 then (Ge t) else if c > 0 then (Ge (CN 0 c t))
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1162
    else (Le (CN 0 (-c) (Neg t))))"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1163
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1164
definition eq :: "int \<Rightarrow> num \<Rightarrow> fm"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1165
where
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1166
  "eq c t = (if c = 0 then (Eq t) else if c > 0 then (Eq (CN 0 c t))
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1167
    else (Eq (CN 0 (-c) (Neg t))))"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1168
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1169
definition neq :: "int \<Rightarrow> num \<Rightarrow> fm"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1170
where
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1171
  "neq c t = (if c = 0 then (NEq t) else if c > 0 then (NEq (CN 0 c t))
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1172
    else (NEq (CN 0 (-c) (Neg t))))"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1173
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60767
diff changeset
  1174
lemma lt: "numnoabs t \<Longrightarrow> Ifm bs (case_prod lt (rsplit0 t)) =
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60767
diff changeset
  1175
  Ifm bs (Lt t) \<and> isrlfm (case_prod lt (rsplit0 t))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1176
  using rsplit0[where bs = "bs" and t="t"]
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1177
  by (auto simp add: lt_def split_def, cases "snd(rsplit0 t)", auto,
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1178
    rename_tac nat a b, case_tac "nat", auto)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1179
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60767
diff changeset
  1180
lemma le: "numnoabs t \<Longrightarrow> Ifm bs (case_prod le (rsplit0 t)) =
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60767
diff changeset
  1181
  Ifm bs (Le t) \<and> isrlfm (case_prod le (rsplit0 t))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1182
  using rsplit0[where bs = "bs" and t="t"]
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1183
  by (auto simp add: le_def split_def, cases "snd(rsplit0 t)", auto,
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1184
    rename_tac nat a b, case_tac "nat", auto)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1185
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60767
diff changeset
  1186
lemma gt: "numnoabs t \<Longrightarrow> Ifm bs (case_prod gt (rsplit0 t)) =
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60767
diff changeset
  1187
  Ifm bs (Gt t) \<and> isrlfm (case_prod gt (rsplit0 t))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1188
  using rsplit0[where bs = "bs" and t="t"]
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1189
  by (auto simp add: gt_def split_def, cases "snd(rsplit0 t)", auto,
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1190
    rename_tac nat a b, case_tac "nat", auto)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1191
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60767
diff changeset
  1192
lemma ge: "numnoabs t \<Longrightarrow> Ifm bs (case_prod ge (rsplit0 t)) =
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60767
diff changeset
  1193
  Ifm bs (Ge t) \<and> isrlfm (case_prod ge (rsplit0 t))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1194
  using rsplit0[where bs = "bs" and t="t"]
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1195
  by (auto simp add: ge_def split_def, cases "snd(rsplit0 t)", auto,
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1196
    rename_tac nat a b, case_tac "nat", auto)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1197
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60767
diff changeset
  1198
lemma eq: "numnoabs t \<Longrightarrow> Ifm bs (case_prod eq (rsplit0 t)) =
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60767
diff changeset
  1199
  Ifm bs (Eq t) \<and> isrlfm (case_prod eq (rsplit0 t))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1200
  using rsplit0[where bs = "bs" and t="t"]
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1201
  by (auto simp add: eq_def split_def, cases "snd(rsplit0 t)", auto,
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1202
    rename_tac nat a b, case_tac "nat", auto)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1203
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60767
diff changeset
  1204
lemma neq: "numnoabs t \<Longrightarrow> Ifm bs (case_prod neq (rsplit0 t)) =
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60767
diff changeset
  1205
  Ifm bs (NEq t) \<and> isrlfm (case_prod neq (rsplit0 t))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1206
  using rsplit0[where bs = "bs" and t="t"]
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1207
  by (auto simp add: neq_def split_def, cases "snd(rsplit0 t)", auto,
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1208
    rename_tac nat a b, case_tac "nat", auto)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1209
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1210
lemma conj_lin: "isrlfm p \<Longrightarrow> isrlfm q \<Longrightarrow> isrlfm (conj p q)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1211
  by (auto simp add: conj_def)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1212
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1213
lemma disj_lin: "isrlfm p \<Longrightarrow> isrlfm q \<Longrightarrow> isrlfm (disj p q)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1214
  by (auto simp add: disj_def)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1215
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1216
fun rlfm :: "fm \<Rightarrow> fm"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1217
where
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1218
  "rlfm (And p q) = conj (rlfm p) (rlfm q)"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1219
| "rlfm (Or p q) = disj (rlfm p) (rlfm q)"
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1220
| "rlfm (Imp p q) = disj (rlfm (Not p)) (rlfm q)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1221
| "rlfm (Iff p q) = disj (conj (rlfm p) (rlfm q)) (conj (rlfm (Not p)) (rlfm (Not q)))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1222
| "rlfm (Lt a) = case_prod lt (rsplit0 a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1223
| "rlfm (Le a) = case_prod le (rsplit0 a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1224
| "rlfm (Gt a) = case_prod gt (rsplit0 a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1225
| "rlfm (Ge a) = case_prod ge (rsplit0 a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1226
| "rlfm (Eq a) = case_prod eq (rsplit0 a)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1227
| "rlfm (NEq a) = case_prod neq (rsplit0 a)"
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1228
| "rlfm (Not (And p q)) = disj (rlfm (Not p)) (rlfm (Not q))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1229
| "rlfm (Not (Or p q)) = conj (rlfm (Not p)) (rlfm (Not q))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1230
| "rlfm (Not (Imp p q)) = conj (rlfm p) (rlfm (Not q))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1231
| "rlfm (Not (Iff p q)) = disj (conj(rlfm p) (rlfm(Not q))) (conj(rlfm(Not p)) (rlfm q))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1232
| "rlfm (Not (Not p)) = rlfm p"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1233
| "rlfm (Not T) = F"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1234
| "rlfm (Not F) = T"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1235
| "rlfm (Not (Lt a)) = rlfm (Ge a)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1236
| "rlfm (Not (Le a)) = rlfm (Gt a)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1237
| "rlfm (Not (Gt a)) = rlfm (Le a)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1238
| "rlfm (Not (Ge a)) = rlfm (Lt a)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1239
| "rlfm (Not (Eq a)) = rlfm (NEq a)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  1240
| "rlfm (Not (NEq a)) = rlfm (Eq a)"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1241
| "rlfm p = p"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1242
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1243
lemma rlfm_I:
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1244
  assumes qfp: "qfree p"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1245
  shows "(Ifm bs (rlfm p) = Ifm bs p) \<and> isrlfm (rlfm p)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1246
  using qfp
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1247
  by (induct p rule: rlfm.induct) (auto simp add: lt le gt ge eq neq conj_lin disj_lin)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1248
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1249
    (* Operations needed for Ferrante and Rackoff *)
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1250
lemma rminusinf_inf:
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1251
  assumes lp: "isrlfm p"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1252
  shows "\<exists>z. \<forall>x < z. Ifm (x#bs) (minusinf p) = Ifm (x#bs) p" (is "\<exists>z. \<forall>x. ?P z x p")
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1253
  using lp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1254
proof (induct p rule: minusinf.induct)
44779
98d597c4193d tuned proofs;
wenzelm
parents: 44013
diff changeset
  1255
  case (1 p q)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1256
  then show ?case
82292
5d91cca0aaf3 Tidied up a few messy proofs
paulson <lp15@cam.ac.uk>
parents: 74397
diff changeset
  1257
    by (fastforce simp: intro: exI [where x= "min _ _"])
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1258
next
44779
98d597c4193d tuned proofs;
wenzelm
parents: 44013
diff changeset
  1259
  case (2 p q)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1260
  then show ?case
82292
5d91cca0aaf3 Tidied up a few messy proofs
paulson <lp15@cam.ac.uk>
parents: 74397
diff changeset
  1261
    by (fastforce simp: intro: exI [where x= "min _ _"])
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1262
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1263
  case (3 c e)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
  1264
  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
  1265
  from 3 have cp: "real_of_int c > 0" by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1266
  fix a
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1267
  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
  1268
  let ?z = "(- ?e) / real_of_int c"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1269
  {
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1270
    fix x
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1271
    assume xz: "x < ?z"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1272
    then have "(real_of_int c * x < - ?e)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1273
      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
  1274
    then have "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
  1275
    then have "real_of_int c * x + ?e \<noteq> 0" by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1276
    with xz have "?P ?z x (Eq (CN 0 c e))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1277
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1278
  }
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1279
  then have "\<forall>x < ?z. ?P ?z x (Eq (CN 0 c e))" by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1280
  then show ?case by blast
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1281
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1282
  case (4 c e)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
  1283
  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
  1284
  from 4 have cp: "real_of_int c > 0" by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1285
  fix a
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1286
  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
  1287
  let ?z = "(- ?e) / real_of_int c"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1288
  {
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1289
    fix x
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1290
    assume xz: "x < ?z"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1291
    then have "(real_of_int c * x < - ?e)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1292
      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
  1293
    then have "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
  1294
    then have "real_of_int c * x + ?e \<noteq> 0" by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1295
    with xz have "?P ?z x (NEq (CN 0 c e))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1296
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1297
  }
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1298
  then have "\<forall>x < ?z. ?P ?z x (NEq (CN 0 c e))" by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1299
  then show ?case by blast
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1300
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1301
  case (5 c e)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
  1302
  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
  1303
  from 5 have cp: "real_of_int c > 0" by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1304
  fix a
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1305
  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
  1306
  let ?z = "(- ?e) / real_of_int c"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1307
  {
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1308
    fix x
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1309
    assume xz: "x < ?z"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1310
    then have "(real_of_int c * x < - ?e)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1311
      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
  1312
    then have "real_of_int c * x + ?e < 0" by arith
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1313
    with xz have "?P ?z x (Lt (CN 0 c e))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1314
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"]  by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1315
  }
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1316
  then have "\<forall>x < ?z. ?P ?z x (Lt (CN 0 c e))" by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1317
  then show ?case by blast
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1318
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1319
  case (6 c e)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
  1320
  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
  1321
  from lp 6 have cp: "real_of_int c > 0" by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1322
  fix a
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1323
  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
  1324
  let ?z = "(- ?e) / real_of_int c"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1325
  {
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1326
    fix x
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1327
    assume xz: "x < ?z"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1328
    then have "(real_of_int c * x < - ?e)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1329
      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
  1330
    then have "real_of_int c * x + ?e < 0" by arith
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1331
    with xz have "?P ?z x (Le (CN 0 c e))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1332
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1333
  }
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1334
  then have "\<forall>x < ?z. ?P ?z x (Le (CN 0 c e))" by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1335
  then show ?case by blast
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1336
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1337
  case (7 c e)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
  1338
  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
  1339
  from 7 have cp: "real_of_int c > 0" by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1340
  fix a
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1341
  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
  1342
  let ?z = "(- ?e) / real_of_int c"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1343
  {
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1344
    fix x
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1345
    assume xz: "x < ?z"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1346
    then have "(real_of_int c * x < - ?e)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1347
      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
  1348
    then have "real_of_int c * x + ?e < 0" by arith
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1349
    with xz have "?P ?z x (Gt (CN 0 c e))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1350
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1351
  }
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1352
  then have "\<forall>x < ?z. ?P ?z x (Gt (CN 0 c e))" by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1353
  then show ?case by blast
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1354
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1355
  case (8 c e)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
  1356
  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
  1357
  from 8 have cp: "real_of_int c > 0" by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1358
  fix a
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1359
  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
  1360
  let ?z = "(- ?e) / real_of_int c"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1361
  {
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1362
    fix x
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1363
    assume xz: "x < ?z"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1364
    then have "(real_of_int c * x < - ?e)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1365
      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
  1366
    then have "real_of_int c * x + ?e < 0" by arith
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1367
    with xz have "?P ?z x (Ge (CN 0 c e))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1368
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1369
  }
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1370
  then have "\<forall>x < ?z. ?P ?z x (Ge (CN 0 c e))" by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1371
  then show ?case by blast
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1372
qed simp_all
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1373
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1374
lemma rplusinf_inf:
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1375
  assumes lp: "isrlfm p"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1376
  shows "\<exists>z. \<forall>x > z. Ifm (x#bs) (plusinf p) = Ifm (x#bs) p" (is "\<exists>z. \<forall>x. ?P z x p")
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1377
using lp
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1378
proof (induct p rule: isrlfm.induct)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1379
  case (1 p q)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1380
  then show ?case
82292
5d91cca0aaf3 Tidied up a few messy proofs
paulson <lp15@cam.ac.uk>
parents: 74397
diff changeset
  1381
    by (fastforce simp: intro: exI [where x= "max _ _"])
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1382
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1383
  case (2 p q)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1384
  then show ?case
82292
5d91cca0aaf3 Tidied up a few messy proofs
paulson <lp15@cam.ac.uk>
parents: 74397
diff changeset
  1385
    by (fastforce simp: intro: exI [where x= "max _ _"])
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1386
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1387
  case (3 c e)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
  1388
  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
  1389
  from 3 have cp: "real_of_int c > 0" by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1390
  fix a
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1391
  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
  1392
  let ?z = "(- ?e) / real_of_int c"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1393
  {
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1394
    fix x
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1395
    assume xz: "x > ?z"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1396
    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
  1397
    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
  1398
    then have "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
  1399
    then have "real_of_int c * x + ?e \<noteq> 0" by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1400
    with xz have "?P ?z x (Eq (CN 0 c e))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1401
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1402
  }
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1403
  then have "\<forall>x > ?z. ?P ?z x (Eq (CN 0 c e))" by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1404
  then show ?case by blast
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1405
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1406
  case (4 c e)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
  1407
  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
  1408
  from 4 have cp: "real_of_int c > 0" by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1409
  fix a
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1410
  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
  1411
  let ?z = "(- ?e) / real_of_int c"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1412
  {
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1413
    fix x
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1414
    assume xz: "x > ?z"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1415
    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
  1416
    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
  1417
    then have "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
  1418
    then have "real_of_int c * x + ?e \<noteq> 0" by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1419
    with xz have "?P ?z x (NEq (CN 0 c e))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1420
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1421
  }
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1422
  then have "\<forall>x > ?z. ?P ?z x (NEq (CN 0 c e))" by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1423
  then show ?case by blast
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1424
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1425
  case (5 c e)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
  1426
  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
  1427
  from 5 have cp: "real_of_int c > 0" by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1428
  fix a
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1429
  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
  1430
  let ?z = "(- ?e) / real_of_int c"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1431
  {
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1432
    fix x
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1433
    assume xz: "x > ?z"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1434
    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
  1435
    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
  1436
    then have "real_of_int c * x + ?e > 0" by arith
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1437
    with xz have "?P ?z x (Lt (CN 0 c e))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1438
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1439
  }
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1440
  then have "\<forall>x > ?z. ?P ?z x (Lt (CN 0 c e))" by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1441
  then show ?case by blast
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1442
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1443
  case (6 c e)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
  1444
  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
  1445
  from 6 have cp: "real_of_int c > 0" by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1446
  fix a
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1447
  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
  1448
  let ?z = "(- ?e) / real_of_int c"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1449
  {
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1450
    fix x
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1451
    assume xz: "x > ?z"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1452
    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
  1453
    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
  1454
    then have "real_of_int c * x + ?e > 0" by arith
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1455
    with xz have "?P ?z x (Le (CN 0 c e))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1456
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1457
  }
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1458
  then have "\<forall>x > ?z. ?P ?z x (Le (CN 0 c e))" by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1459
  then show ?case by blast
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1460
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1461
  case (7 c e)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
  1462
  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
  1463
  from 7 have cp: "real_of_int c > 0" by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1464
  fix a
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1465
  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
  1466
  let ?z = "(- ?e) / real_of_int c"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1467
  {
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1468
    fix x
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1469
    assume xz: "x > ?z"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1470
    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
  1471
    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
  1472
    then have "real_of_int c * x + ?e > 0" by arith
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1473
    with xz have "?P ?z x (Gt (CN 0 c e))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1474
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1475
  }
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1476
  then have "\<forall>x > ?z. ?P ?z x (Gt (CN 0 c e))" by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1477
  then show ?case by blast
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1478
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1479
  case (8 c e)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
  1480
  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
  1481
  from 8 have cp: "real_of_int c > 0" by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1482
  fix a
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1483
  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
  1484
  let ?z = "(- ?e) / real_of_int c"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1485
  {
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1486
    fix x
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1487
    assume xz: "x > ?z"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1488
    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
  1489
    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
  1490
    then have "real_of_int c * x + ?e > 0" by arith
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1491
    with xz have "?P ?z x (Ge (CN 0 c e))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1492
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1493
  }
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1494
  then have "\<forall>x > ?z. ?P ?z x (Ge (CN 0 c e))" by simp
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1495
  then show ?case by blast
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1496
qed simp_all
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1497
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1498
lemma rminusinf_bound0:
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1499
  assumes lp: "isrlfm p"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1500
  shows "bound0 (minusinf p)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1501
  using lp by (induct p rule: minusinf.induct) simp_all
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1502
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1503
lemma rplusinf_bound0:
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1504
  assumes lp: "isrlfm p"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1505
  shows "bound0 (plusinf p)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1506
  using lp by (induct p rule: plusinf.induct) simp_all
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1507
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1508
lemma rminusinf_ex:
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1509
  assumes lp: "isrlfm p"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1510
    and ex: "Ifm (a#bs) (minusinf p)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1511
  shows "\<exists>x. Ifm (x#bs) p"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1512
proof -
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1513
  from bound0_I [OF rminusinf_bound0[OF lp], where b="a" and bs ="bs"] ex
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1514
  have th: "\<forall>x. Ifm (x#bs) (minusinf p)" by auto
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1515
  from rminusinf_inf[OF lp, where bs="bs"]
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1516
  obtain z where z_def: "\<forall>x<z. Ifm (x # bs) (minusinf p) = Ifm (x # bs) p" by blast
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1517
  from th have "Ifm ((z - 1) # bs) (minusinf p)" by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1518
  moreover have "z - 1 < z" by simp
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1519
  ultimately show ?thesis using z_def by auto
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1520
qed
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1521
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1522
lemma rplusinf_ex:
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1523
  assumes lp: "isrlfm p"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1524
    and ex: "Ifm (a # bs) (plusinf p)"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1525
  shows "\<exists>x. Ifm (x # bs) p"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1526
proof -
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1527
  from bound0_I [OF rplusinf_bound0[OF lp], where b="a" and bs ="bs"] ex
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1528
  have th: "\<forall>x. Ifm (x # bs) (plusinf p)" by auto
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1529
  from rplusinf_inf[OF lp, where bs="bs"]
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1530
  obtain z where z_def: "\<forall>x>z. Ifm (x # bs) (plusinf p) = Ifm (x # bs) p" by blast
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1531
  from th have "Ifm ((z + 1) # bs) (plusinf p)" by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1532
  moreover have "z + 1 > z" by simp
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1533
  ultimately show ?thesis using z_def by auto
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1534
qed
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1535
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1536
fun uset :: "fm \<Rightarrow> (num \<times> int) list"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1537
where
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1538
  "uset (And p q) = (uset p @ uset q)"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1539
| "uset (Or p q) = (uset p @ uset q)"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1540
| "uset (Eq  (CN 0 c e)) = [(Neg e,c)]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1541
| "uset (NEq (CN 0 c e)) = [(Neg e,c)]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1542
| "uset (Lt  (CN 0 c e)) = [(Neg e,c)]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1543
| "uset (Le  (CN 0 c e)) = [(Neg e,c)]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1544
| "uset (Gt  (CN 0 c e)) = [(Neg e,c)]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1545
| "uset (Ge  (CN 0 c e)) = [(Neg e,c)]"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1546
| "uset p = []"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1547
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1548
fun usubst :: "fm \<Rightarrow> num \<times> int \<Rightarrow> fm"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1549
where
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1550
  "usubst (And p q) = (\<lambda>(t,n). And (usubst p (t,n)) (usubst q (t,n)))"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1551
| "usubst (Or p q) = (\<lambda>(t,n). Or (usubst p (t,n)) (usubst q (t,n)))"
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1552
| "usubst (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: 66453
diff changeset
  1553
| "usubst (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: 66453
diff changeset
  1554
| "usubst (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: 66453
diff changeset
  1555
| "usubst (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: 66453
diff changeset
  1556
| "usubst (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: 66453
diff changeset
  1557
| "usubst (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: 66453
diff changeset
  1558
| "usubst p = (\<lambda>(t, n). p)"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1559
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1560
lemma usubst_I:
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1561
  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
  1562
    and np: "real_of_int n > 0"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1563
    and nbt: "numbound0 t"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1564
  shows "(Ifm (x # bs) (usubst p (t,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
  1565
    Ifm (((Inum (x # bs) t) / (real_of_int n)) # bs) p) \<and> bound0 (usubst p (t, n))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1566
  (is "(?I x (usubst p (t, n)) = ?I ?u p) \<and> ?B p"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1567
   is "(_ = ?I (?t/?n) p) \<and> _"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1568
   is "(_ = ?I (?N x t /_) p) \<and> _")
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1569
  using lp
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1570
proof (induct p rule: usubst.induct)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1571
  case (5 c e)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1572
  with assms 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
  1573
  have "?I ?u (Lt (CN 0 c e)) \<longleftrightarrow> real_of_int c * (?t / ?n) + ?N x e < 0"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1574
    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
  1575
  also have "\<dots> \<longleftrightarrow> ?n * (real_of_int c * (?t / ?n)) + ?n*(?N x 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
  1576
    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: 61945
diff changeset
  1577
      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
  1578
  also have "\<dots> \<longleftrightarrow> real_of_int c * ?t + ?n * (?N x e) < 0" using np by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1579
  finally show ?case using nbt nb by (simp add: algebra_simps)
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1580
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1581
  case (6 c e)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1582
  with assms 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
  1583
  have "?I ?u (Le (CN 0 c e)) \<longleftrightarrow> real_of_int c * (?t / ?n) + ?N x e \<le> 0"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1584
    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
  1585
  also have "\<dots> = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x 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
  1586
    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: 61945
diff changeset
  1587
      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
  1588
  also have "\<dots> = (real_of_int c *?t + ?n* (?N x e) \<le> 0)" using np by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1589
  finally show ?case using nbt nb by (simp add: algebra_simps)
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1590
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1591
  case (7 c e)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1592
  with assms 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
  1593
  have "?I ?u (Gt (CN 0 c e)) \<longleftrightarrow> real_of_int c *(?t / ?n) + ?N x e > 0"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1594
    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
  1595
  also have "\<dots> \<longleftrightarrow> ?n * (real_of_int c * (?t / ?n)) + ?n * ?N x 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
  1596
    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: 61945
diff changeset
  1597
      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
  1598
  also have "\<dots> \<longleftrightarrow> real_of_int c * ?t + ?n * ?N x e > 0" using np by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1599
  finally show ?case using nbt nb by (simp add: algebra_simps)
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1600
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1601
  case (8 c e)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1602
  with assms 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
  1603
  have "?I ?u (Ge (CN 0 c e)) \<longleftrightarrow> real_of_int c * (?t / ?n) + ?N x e \<ge> 0"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1604
    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
  1605
  also have "\<dots> \<longleftrightarrow> ?n * (real_of_int c * (?t / ?n)) + ?n * ?N x 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
  1606
    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: 61945
diff changeset
  1607
      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
  1608
  also have "\<dots> \<longleftrightarrow> real_of_int c * ?t + ?n * ?N x e \<ge> 0" using np by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1609
  finally show ?case using nbt nb by (simp add: algebra_simps)
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1610
next
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1611
  case (3 c e)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1612
  with assms 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
  1613
  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
  1614
  have "?I ?u (Eq (CN 0 c e)) \<longleftrightarrow> real_of_int c * (?t / ?n) + ?N x e = 0"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1615
    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
  1616
  also have "\<dots> \<longleftrightarrow> ?n * (real_of_int c * (?t / ?n)) + ?n * ?N x 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
  1617
    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: 61945
diff changeset
  1618
      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
  1619
  also have "\<dots> \<longleftrightarrow> real_of_int c * ?t + ?n * ?N x e = 0" using np by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1620
  finally show ?case using nbt nb by (simp add: algebra_simps)
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1621
next
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
  1622
  case (4 c e) with assms 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
  1623
  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
  1624
  have "?I ?u (NEq (CN 0 c e)) \<longleftrightarrow> real_of_int c * (?t / ?n) + ?N x e \<noteq> 0"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1625
    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
  1626
  also have "\<dots> \<longleftrightarrow> ?n * (real_of_int c * (?t / ?n)) + ?n * ?N x 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
  1627
    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: 61945
diff changeset
  1628
      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
  1629
  also have "\<dots> \<longleftrightarrow> real_of_int c * ?t + ?n * ?N x e \<noteq> 0" using np by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1630
  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
  1631
qed(simp_all add: nbt numbound0_I[where bs ="bs" and b="(Inum (x#bs) t)/ real_of_int n" and b'="x"])
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1632
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1633
lemma uset_l:
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1634
  assumes lp: "isrlfm p"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1635
  shows "\<forall>(t,k) \<in> set (uset p). numbound0 t \<and> k > 0"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1636
  using lp by (induct p rule: uset.induct) auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1637
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1638
lemma rminusinf_uset:
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1639
  assumes lp: "isrlfm p"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1640
    and nmi: "\<not> (Ifm (a # bs) (minusinf p))" (is "\<not> (Ifm (a # bs) (?M p))")
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1641
    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
  1642
  shows "\<exists>(s,m) \<in> set (uset p). x \<ge> Inum (a#bs) s / real_of_int m"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1643
    (is "\<exists>(s,m) \<in> ?U p. x \<ge> ?N a s / real_of_int m")
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1644
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
  1645
  have "\<exists>(s,m) \<in> set (uset p). real_of_int m * x \<ge> Inum (a#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
  1646
    (is "\<exists>(s,m) \<in> ?U p. real_of_int m *x \<ge> ?N a s")
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1647
    using lp nmi ex
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1648
    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
  1649
  then obtain s m where smU: "(s,m) \<in> set (uset p)" and mx: "real_of_int m * x \<ge> ?N a s"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1650
    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
  1651
  from uset_l[OF lp] smU have mp: "real_of_int m > 0"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1652
    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
  1653
  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: 56544
diff changeset
  1654
    by (auto simp add: mult.commute)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1655
  then show ?thesis
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1656
    using smU by auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1657
qed
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1658
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1659
lemma rplusinf_uset:
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1660
  assumes lp: "isrlfm p"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1661
    and nmi: "\<not> (Ifm (a # bs) (plusinf p))" (is "\<not> (Ifm (a # bs) (?M p))")
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1662
    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
  1663
  shows "\<exists>(s,m) \<in> set (uset p). x \<le> Inum (a#bs) s / real_of_int m"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1664
    (is "\<exists>(s,m) \<in> ?U p. x \<le> ?N a s / real_of_int m")
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1665
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
  1666
  have "\<exists>(s,m) \<in> set (uset p). real_of_int m * x \<le> Inum (a#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
  1667
    (is "\<exists>(s,m) \<in> ?U p. real_of_int m *x \<le> ?N a s")
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1668
    using lp nmi ex
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1669
    by (induct p rule: minusinf.induct)
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1670
      (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
  1671
  then obtain s m where smU: "(s,m) \<in> set (uset p)" and mx: "real_of_int m * x \<le> ?N a s"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1672
    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
  1673
  from uset_l[OF lp] smU have mp: "real_of_int m > 0"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1674
    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
  1675
  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: 56544
diff changeset
  1676
    by (auto simp add: mult.commute)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1677
  then show ?thesis
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1678
    using smU by auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1679
qed
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1680
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1681
lemma lin_dense:
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1682
  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
  1683
    and noS: "\<forall>t. l < t \<and> t< u \<longrightarrow> t \<notin> (\<lambda>(t,n). Inum (x#bs) t / real_of_int n) ` set (uset 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
  1684
      (is "\<forall>t. _ \<and> _ \<longrightarrow> t \<notin> (\<lambda>(t,n). ?N x t / real_of_int n ) ` (?U p)")
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1685
    and lx: "l < x"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1686
    and xu:"x < u"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1687
    and px:" Ifm (x#bs) p"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1688
    and ly: "l < y" and yu: "y < u"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1689
  shows "Ifm (y#bs) p"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1690
  using lp px noS
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1691
proof (induct p rule: isrlfm.induct)
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1692
  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
  1693
  then have cp: "real_of_int c > 0" and nb: "numbound0 e"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1694
    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
  1695
  from 5 have "x * real_of_int c + ?N x e < 0"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1696
    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
  1697
  then have pxc: "x < (- ?N x e) / real_of_int c"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
  1698
    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
  1699
  from 5 have noSc:"\<forall>t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real_of_int c"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1700
    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
  1701
  with ly yu have yne: "y \<noteq> - ?N x e / real_of_int c"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1702
    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
  1703
  then consider "y < (-?N x e)/ real_of_int c" | "y > (- ?N x e) / real_of_int c"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1704
    by atomize_elim auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1705
  then show ?case
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1706
  proof cases
60767
ad5b4771fc19 tuned proofs;
wenzelm
parents: 60711
diff changeset
  1707
    case 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
  1708
    then have "y * real_of_int c < - ?N x e"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1709
      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
  1710
    then have "real_of_int c * y + ?N x e < 0"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1711
      by (simp add: algebra_simps)
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1712
    then show ?thesis
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1713
      using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1714
  next
60767
ad5b4771fc19 tuned proofs;
wenzelm
parents: 60711
diff changeset
  1715
    case 2
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1716
    with yu have eu: "u > (- ?N x e) / real_of_int c"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1717
      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
  1718
    with noSc ly yu have "(- ?N x e) / real_of_int c \<le> l"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1719
      by (cases "(- ?N x e) / real_of_int c > l") auto
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1720
    with lx pxc have False
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1721
      by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1722
    then show ?thesis ..
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1723
  qed
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1724
next
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1725
  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
  1726
  then have cp: "real_of_int c > 0" and nb: "numbound0 e"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1727
    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
  1728
  from 6 have "x * real_of_int c + ?N x e \<le> 0"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1729
    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
  1730
  then have pxc: "x \<le> (- ?N x e) / real_of_int c"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1731
    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
  1732
  from 6 have noSc:"\<forall>t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real_of_int c"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1733
    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
  1734
  with ly yu have yne: "y \<noteq> - ?N x e / real_of_int c"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1735
    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
  1736
  then consider "y < (- ?N x e) / real_of_int c" | "y > (-?N x e) / real_of_int c"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1737
    by atomize_elim auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1738
  then show ?case
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1739
  proof cases
60767
ad5b4771fc19 tuned proofs;
wenzelm
parents: 60711
diff changeset
  1740
    case 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
  1741
    then have "y * real_of_int c < - ?N x e"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
  1742
      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
  1743
    then have "real_of_int c * y + ?N x e < 0"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1744
      by (simp add: algebra_simps)
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1745
    then show ?thesis
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1746
      using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1747
  next
60767
ad5b4771fc19 tuned proofs;
wenzelm
parents: 60711
diff changeset
  1748
    case 2
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1749
    with yu have eu: "u > (- ?N x e) / real_of_int c"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1750
      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
  1751
    with noSc ly yu have "(- ?N x e) / real_of_int c \<le> l"
77b453bd616f Coercion "real" now has type nat => real only 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
      by (cases "(- ?N x e) / real_of_int c > l") auto
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1753
    with lx pxc have False
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1754
      by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1755
    then show ?thesis ..
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1756
  qed
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1757
next
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1758
  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
  1759
  then have cp: "real_of_int c > 0" and nb: "numbound0 e"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1760
    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
  1761
  from 7 have "x * real_of_int c + ?N x e > 0"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1762
    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
  1763
  then have pxc: "x > (- ?N x e) / real_of_int c"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
  1764
    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
  1765
  from 7 have noSc: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real_of_int c"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1766
    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
  1767
  with ly yu have yne: "y \<noteq> - ?N x e / real_of_int c"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1768
    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
  then consider "y > (- ?N x e) / real_of_int c" | "y < (-?N x e) / real_of_int c"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1770
    by atomize_elim auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1771
  then show ?case
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1772
  proof cases
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1773
    case 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
  1774
    then have "y * real_of_int c > - ?N x e"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1775
      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
  1776
    then have "real_of_int c * y + ?N x e > 0"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1777
      by (simp add: algebra_simps)
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1778
    then show ?thesis
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1779
      using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1780
  next
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1781
    case 2
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1782
    with ly have eu: "l < (- ?N x e) / real_of_int c"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1783
      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
  1784
    with noSc ly yu have "(- ?N x e) / real_of_int c \<ge> u"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1785
      by (cases "(- ?N x e) / real_of_int c > l") auto
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1786
    with xu pxc have False by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1787
    then show ?thesis ..
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1788
  qed
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1789
next
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1790
  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
  1791
  then have cp: "real_of_int c > 0" and nb: "numbound0 e"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1792
    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
  1793
  from 8 have "x * real_of_int c + ?N x e \<ge> 0"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1794
    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
  1795
  then have pxc: "x \<ge> (- ?N x e) / real_of_int c"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1796
    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
  1797
  from 8 have noSc:"\<forall>t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real_of_int c"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1798
    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
  1799
  with ly yu have yne: "y \<noteq> - ?N x e / real_of_int c"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1800
    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
  1801
  then consider "y > (- ?N x e) / real_of_int c" | "y < (-?N x e) / real_of_int c"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1802
    by atomize_elim auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1803
  then show ?case
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1804
  proof cases
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1805
    case 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
  1806
    then have "y * real_of_int c > - ?N x e"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
  1807
      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
  1808
    then have "real_of_int c * y + ?N x e > 0" by (simp add: algebra_simps)
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1809
    then show ?thesis
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1810
      using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1811
  next
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1812
    case 2
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1813
    with ly have eu: "l < (- ?N x e) / real_of_int c"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1814
      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
  1815
    with noSc ly yu have "(- ?N x e) / real_of_int c \<ge> u"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1816
      by (cases "(- ?N x e) / real_of_int c > l") auto
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1817
    with xu pxc have False
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1818
      by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1819
    then show ?thesis ..
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1820
  qed
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1821
next
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1822
  case (3 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
  1823
  then have cp: "real_of_int c > 0" and nb: "numbound0 e"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1824
    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
  1825
  from cp have cnz: "real_of_int c \<noteq> 0"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1826
    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
  1827
  from 3 have "x * real_of_int c + ?N x e = 0"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1828
    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
  1829
  then have pxc: "x = (- ?N x e) / real_of_int c"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
  1830
    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
  1831
  from 3 have noSc:"\<forall>t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real_of_int c"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1832
    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
  1833
  with lx xu have yne: "x \<noteq> - ?N x e / real_of_int c"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1834
    by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1835
  with pxc show ?case
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1836
    by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1837
next
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1838
  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
  1839
  then have cp: "real_of_int c > 0" and nb: "numbound0 e"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1840
    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
  1841
  from cp have cnz: "real_of_int c \<noteq> 0"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1842
    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
  1843
  from 4 have noSc:"\<forall>t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real_of_int c"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1844
    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
  1845
  with ly yu have yne: "y \<noteq> - ?N x e / real_of_int c"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1846
    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
  1847
  then have "y* real_of_int c \<noteq> -?N x e"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
  1848
    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
  1849
  then have "y* real_of_int c + ?N x e \<noteq> 0"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1850
    by (simp add: algebra_simps)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1851
  then show ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"]
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
  1852
    by (simp add: algebra_simps)
41842
d8f76db6a207 added simp lemma nth_Cons_pos to List
nipkow
parents: 41838
diff changeset
  1853
qed (auto simp add: numbound0_I[where bs="bs" and b="y" and b'="x"])
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1854
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1855
lemma finite_set_intervals:
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1856
  fixes x :: real
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1857
  assumes px: "P x"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1858
    and lx: "l \<le> x"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1859
    and xu: "x \<le> u"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1860
    and linS: "l\<in> S"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1861
    and uinS: "u \<in> S"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1862
    and fS: "finite S"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1863
    and lS: "\<forall>x\<in> S. l \<le> x"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1864
    and Su: "\<forall>x\<in> S. x \<le> u"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1865
  shows "\<exists>a \<in> S. \<exists>b \<in> S. (\<forall>y. a < y \<and> y < b \<longrightarrow> y \<notin> S) \<and> a \<le> x \<and> x \<le> b \<and> P x"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1866
proof -
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1867
  let ?Mx = "{y. y\<in> S \<and> y \<le> x}"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1868
  let ?xM = "{y. y\<in> S \<and> x \<le> y}"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1869
  let ?a = "Max ?Mx"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1870
  let ?b = "Min ?xM"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1871
  have MxS: "?Mx \<subseteq> S"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1872
    by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1873
  then have fMx: "finite ?Mx"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1874
    using fS finite_subset by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1875
  from lx linS have linMx: "l \<in> ?Mx"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1876
    by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1877
  then have Mxne: "?Mx \<noteq> {}"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1878
    by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1879
  have xMS: "?xM \<subseteq> S"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1880
    by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1881
  then have fxM: "finite ?xM"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1882
    using fS finite_subset by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1883
  from xu uinS have linxM: "u \<in> ?xM"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1884
    by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1885
  then have xMne: "?xM \<noteq> {}"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1886
    by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1887
  have ax:"?a \<le> x"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1888
    using Mxne fMx by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1889
  have xb:"x \<le> ?b"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1890
    using xMne fxM by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1891
  have "?a \<in> ?Mx"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1892
    using Max_in[OF fMx Mxne] by simp
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1893
  then have ainS: "?a \<in> S"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1894
    using MxS by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1895
  have "?b \<in> ?xM"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1896
    using Min_in[OF fxM xMne] by simp
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1897
  then have binS: "?b \<in> S"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1898
    using xMS by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1899
  have noy: "\<forall>y. ?a < y \<and> y < ?b \<longrightarrow> y \<notin> S"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1900
  proof clarsimp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1901
    fix y
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1902
    assume ay: "?a < y" and yb: "y < ?b" and yS: "y \<in> S"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1903
    from yS consider "y \<in> ?Mx" | "y \<in> ?xM"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1904
      by atomize_elim auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1905
    then show False
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1906
    proof cases
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1907
      case 1
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1908
      then have "y \<le> ?a"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1909
        using Mxne fMx by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1910
      with ay show ?thesis by simp
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1911
    next
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1912
      case 2
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1913
      then have "y \<ge> ?b"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1914
        using xMne fxM by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1915
      with yb show ?thesis by simp
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1916
    qed
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1917
  qed
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1918
  from ainS binS noy ax xb px show ?thesis
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1919
    by blast
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1920
qed
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1921
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1922
lemma rinf_uset:
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1923
  assumes lp: "isrlfm p"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1924
    and nmi: "\<not> (Ifm (x # bs) (minusinf p))"  (is "\<not> (Ifm (x # bs) (?M p))")
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1925
    and npi: "\<not> (Ifm (x # bs) (plusinf p))"  (is "\<not> (Ifm (x # bs) (?P p))")
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1926
    and ex: "\<exists>x. Ifm (x # bs) p"  (is "\<exists>x. ?I x p")
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1927
  shows "\<exists>(l,n) \<in> set (uset p). \<exists>(s,m) \<in> set (uset 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
  1928
    ?I ((Inum (x#bs) l / real_of_int n + Inum (x#bs) s / real_of_int m) / 2) p"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1929
proof -
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1930
  let ?N = "\<lambda>x t. Inum (x # bs) t"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1931
  let ?U = "set (uset p)"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1932
  from ex obtain a where pa: "?I a p"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1933
    by blast
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1934
  from bound0_I[OF rminusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] nmi
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1935
  have nmi': "\<not> (?I a (?M p))"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1936
    by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1937
  from bound0_I[OF rplusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] npi
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1938
  have npi': "\<not> (?I a (?P p))"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1939
    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
  1940
  have "\<exists>(l,n) \<in> set (uset p). \<exists>(s,m) \<in> set (uset p). ?I ((?N a l/real_of_int n + ?N a s /real_of_int m) / 2) p"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1941
  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
  1942
    let ?M = "(\<lambda>(t,c). ?N a t / real_of_int c) ` ?U"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1943
    have fM: "finite ?M"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1944
      by auto
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1945
    from rminusinf_uset[OF lp nmi pa] rplusinf_uset[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
  1946
    have "\<exists>(l,n) \<in> set (uset p). \<exists>(s,m) \<in> set (uset p). a \<le> ?N x l / real_of_int n \<and> a \<ge> ?N x s / real_of_int m"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1947
      by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1948
    then obtain "t" "n" "s" "m"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1949
      where tnU: "(t,n) \<in> ?U"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1950
        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
  1951
        and xs1: "a \<le> ?N x s / real_of_int m"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1952
        and tx1: "a \<ge> ?N x t / real_of_int n"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1953
      by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1954
    from uset_l[OF lp] tnU smU numbound0_I[where bs="bs" and b="x" and b'="a"] xs1 tx1
61609
77b453bd616f Coercion "real" now has type nat => real only 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
    have xs: "a \<le> ?N a s / real_of_int m" and tx: "a \<ge> ?N a t / real_of_int n"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1956
      by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1957
    from tnU have Mne: "?M \<noteq> {}"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1958
      by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1959
    then have Une: "?U \<noteq> {}"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1960
      by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1961
    let ?l = "Min ?M"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  1962
    let ?u = "Max ?M"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1963
    have linM: "?l \<in> ?M"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1964
      using fM Mne by simp
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1965
    have uinM: "?u \<in> ?M"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1966
      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
  1967
    have tnM: "?N a t / real_of_int n \<in> ?M"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1968
      using tnU 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
  1969
    have smM: "?N a s / real_of_int m \<in> ?M"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1970
      using smU by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1971
    have lM: "\<forall>t\<in> ?M. ?l \<le> t"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1972
      using Mne fM by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1973
    have Mu: "\<forall>t\<in> ?M. t \<le> ?u"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1974
      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
  1975
    have "?l \<le> ?N a t / real_of_int n"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1976
      using tnM Mne by simp
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1977
    then have lx: "?l \<le> a"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1978
      using tx 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
  1979
    have "?N a s / real_of_int m \<le> ?u"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1980
      using smM Mne by simp
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1981
    then have xu: "a \<le> ?u"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1982
      using xs by simp
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  1983
    from finite_set_intervals2[where P="\<lambda>x. ?I x p",OF pa lx xu linM uinM fM lM Mu]
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1984
    consider u where "u \<in> ?M" "?I u p"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1985
      | t1 t2 where "t1 \<in> ?M" "t2 \<in> ?M" "\<forall>y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M" "t1 < a" "a < t2" "?I a p"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1986
      by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1987
    then show ?thesis
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1988
    proof cases
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1989
      case 1
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1990
      note um = \<open>u \<in> ?M\<close> and pu = \<open>?I u p\<close>
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1991
      then have "\<exists>(tu,nu) \<in> ?U. u = ?N a tu / real_of_int nu"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1992
        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
  1993
      then obtain tu nu where tuU: "(tu, nu) \<in> ?U" and tuu: "u= ?N a tu / real_of_int nu"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32642
diff changeset
  1994
        by blast
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1995
      have "(u + u) / 2 = u"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1996
        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
  1997
      with pu tuu have "?I (((?N a tu / real_of_int nu) + (?N a tu / real_of_int nu)) / 2) p"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1998
        by simp
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  1999
      with tuU show ?thesis by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2000
    next
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2001
      case 2
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2002
      note t1M = \<open>t1 \<in> ?M\<close> and t2M = \<open>t2\<in> ?M\<close>
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2003
        and noM = \<open>\<forall>y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M\<close>
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2004
        and t1x = \<open>t1 < a\<close> and xt2 = \<open>a < t2\<close> and px = \<open>?I a p\<close>
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2005
      from t1M have "\<exists>(t1u,t1n) \<in> ?U. t1 = ?N a t1u / real_of_int t1n"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2006
        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
  2007
      then obtain t1u t1n where t1uU: "(t1u, t1n) \<in> ?U" and t1u: "t1 = ?N a t1u / real_of_int t1n"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2008
        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
  2009
      from t2M have "\<exists>(t2u,t2n) \<in> ?U. t2 = ?N a t2u / real_of_int t2n"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2010
        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
  2011
      then obtain t2u t2n where t2uU: "(t2u, t2n) \<in> ?U" and t2u: "t2 = ?N a t2u / real_of_int t2n"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2012
        by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2013
      from t1x xt2 have t1t2: "t1 < t2"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2014
        by simp
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2015
      let ?u = "(t1 + t2) / 2"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2016
      from less_half_sum[OF t1t2] gt_half_sum[OF t1t2] have t1lu: "t1 < ?u" and ut2: "?u < t2"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2017
        by auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2018
      from lin_dense[OF lp noM t1x xt2 px t1lu ut2] have "?I ?u p" .
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2019
      with t1uU t2uU t1u t2u show ?thesis
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2020
        by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2021
    qed
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2022
  qed
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2023
  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
  2024
    and pu: "?I ((?N a l / real_of_int n + ?N a s / real_of_int m) / 2) p"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2025
    by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2026
  from lnU smU uset_l[OF lp] have nbl: "numbound0 l" and nbs: "numbound0 s"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2027
    by auto
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  2028
  from numbound0_I[OF nbl, where bs="bs" and b="a" and b'="x"]
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2029
    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
  2030
  have "?I ((?N x l / real_of_int n + ?N x s / real_of_int m) / 2) p"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2031
    by simp
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2032
  with lnU smU show ?thesis
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2033
    by auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2034
qed
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2035
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2036
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2037
    (* The Ferrante - Rackoff Theorem *)
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2038
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  2039
theorem fr_eq:
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2040
  assumes lp: "isrlfm p"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2041
  shows "(\<exists>x. Ifm (x#bs) p) \<longleftrightarrow>
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2042
    Ifm (x # bs) (minusinf p) \<or> Ifm (x # bs) (plusinf p) \<or>
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2043
      (\<exists>(t,n) \<in> set (uset p). \<exists>(s,m) \<in> set (uset 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
  2044
        Ifm ((((Inum (x # bs) t) / real_of_int n + (Inum (x # bs) s) / real_of_int m) / 2) # bs) p)"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2045
  (is "(\<exists>x. ?I x p) \<longleftrightarrow> (?M \<or> ?P \<or> ?F)" is "?E = ?D")
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2046
proof
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  2047
  assume px: "\<exists>x. ?I x p"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2048
  consider "?M \<or> ?P" | "\<not> ?M" "\<not> ?P" by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2049
  then show ?D
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2050
  proof cases
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2051
    case 1
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2052
    then show ?thesis by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2053
  next
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2054
    case 2
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2055
    from rinf_uset[OF lp this] have ?F
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2056
      using px by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2057
    then show ?thesis by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2058
  qed
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2059
next
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2060
  assume ?D
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2061
  then consider ?M | ?P | ?F by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2062
  then show ?E
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2063
  proof cases
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2064
    case 1
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2065
    from rminusinf_ex[OF lp this] show ?thesis .
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2066
  next
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2067
    case 2
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2068
    from rplusinf_ex[OF lp this] show ?thesis .
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2069
  next
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2070
    case 3
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2071
    then show ?thesis by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2072
  qed
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2073
qed
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2074
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2075
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  2076
lemma fr_equsubst:
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2077
  assumes lp: "isrlfm p"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2078
  shows "(\<exists>x. Ifm (x # bs) p) \<longleftrightarrow>
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2079
    (Ifm (x # bs) (minusinf p) \<or> Ifm (x # bs) (plusinf p) \<or>
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2080
      (\<exists>(t,k) \<in> set (uset p). \<exists>(s,l) \<in> set (uset p).
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2081
        Ifm (x#bs) (usubst p (Add (Mul l t) (Mul k s), 2 * k * l))))"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2082
  (is "(\<exists>x. ?I x p) \<longleftrightarrow> ?M \<or> ?P \<or> ?F" is "?E = ?D")
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2083
proof
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  2084
  assume px: "\<exists>x. ?I x p"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2085
  consider "?M \<or> ?P" | "\<not> ?M" "\<not> ?P" by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2086
  then show ?D
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2087
  proof cases
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2088
    case 1
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2089
    then show ?thesis by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2090
  next
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2091
    case 2
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2092
    let ?f = "\<lambda>(t,n). Inum (x # bs) t / real_of_int n"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2093
    let ?N = "\<lambda>t. Inum (x # bs) t"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2094
    {
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2095
      fix t n s m
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2096
      assume "(t, n) \<in> set (uset p)" and "(s, m) \<in> set (uset p)"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2097
      with uset_l[OF lp] have 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
  2098
        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: 32642
diff changeset
  2099
        by auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2100
      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
  2101
      from np mp have mnp: "real_of_int (2 * n * m) > 0"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2102
        by (simp add: mult.commute)
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2103
      from tnb snb have st_nb: "numbound0 ?st"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2104
        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
  2105
      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: 32642
diff changeset
  2106
        using mnp mp np by (simp add: algebra_simps add_divide_distrib)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  2107
      from usubst_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
  2108
      have "?I x (usubst p (?st, 2 * n * m)) = ?I ((?N t / real_of_int n + ?N s / real_of_int m) / 2) p"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2109
        by (simp only: st[symmetric])
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2110
    }
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2111
    with rinf_uset[OF lp 2 px] have ?F
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2112
      by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2113
    then show ?thesis
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2114
      by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2115
  qed
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2116
next
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2117
  assume ?D
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2118
  then consider ?M | ?P | t k s l where "(t, k) \<in> set (uset p)" "(s, l) \<in> set (uset p)"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2119
    "?I x (usubst p (Add (Mul l t) (Mul k s), 2 * k * l))"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2120
    by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2121
  then show ?E
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2122
  proof cases
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2123
    case 1
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2124
    from rminusinf_ex[OF lp this] show ?thesis .
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2125
  next
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2126
    case 2
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2127
    from rplusinf_ex[OF lp this] show ?thesis .
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2128
  next
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2129
    case 3
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2130
    with uset_l[OF lp] have tnb: "numbound0 t" and np: "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
  2131
      and snb: "numbound0 s" and mp: "real_of_int l > 0"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2132
      by auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2133
    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
  2134
    from np mp have mnp: "real_of_int (2 * k * l) > 0"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2135
      by (simp add: mult.commute)
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2136
    from tnb snb have st_nb: "numbound0 ?st"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2137
      by simp
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2138
    from usubst_I[OF lp mnp st_nb, where bs="bs"]
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2139
      \<open>?I x (usubst p (Add (Mul l t) (Mul k s), 2 * k * l))\<close> show ?thesis
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2140
      by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2141
  qed
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2142
qed
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2143
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2144
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2145
    (* Implement the right hand side of Ferrante and Rackoff's Theorem. *)
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2146
definition ferrack :: "fm \<Rightarrow> fm"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2147
where
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2148
  "ferrack p =
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2149
   (let
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2150
      p' = rlfm (simpfm p);
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2151
      mp = minusinf p';
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2152
      pp = plusinf p'
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2153
    in
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2154
      if mp = T \<or> pp = T then T
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2155
      else
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2156
       (let U = remdups (map simp_num_pair
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2157
         (map (\<lambda>((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2 * n * m))
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2158
               (alluopairs (uset p'))))
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2159
        in decr (disj mp (disj pp (evaldjf (simpfm \<circ> usubst p') U)))))"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2160
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2161
lemma uset_cong_aux:
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2162
  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
  2163
  shows "((\<lambda>(t,n). Inum (x # bs) t / real_of_int n) `
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2164
    (set (map (\<lambda>((t,n),(s,m)). (Add (Mul m t) (Mul n s), 2 * n * m)) (alluopairs 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
  2165
    ((\<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))"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2166
  (is "?lhs = ?rhs")
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2167
proof auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2168
  fix t n s m
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2169
  assume "((t, n), (s, m)) \<in> set (alluopairs U)"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2170
  then have th: "((t, n), (s, m)) \<in> set U \<times> set U"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2171
    using alluopairs_set1[where xs="U"] by blast
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2172
  let ?N = "\<lambda>t. Inum (x # bs) t"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2173
  let ?st = "Add (Mul m t) (Mul n s)"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2174
  from Ul th have mnz: "m \<noteq> 0"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2175
    by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2176
  from Ul th have nnz: "n \<noteq> 0"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2177
    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
  2178
  have st: "(?N t / real_of_int n + ?N s / real_of_int m) / 2 = ?N ?st / real_of_int (2 * n * m)"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2179
    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
  2180
  then show "(real_of_int m *  Inum (x # bs) t + real_of_int n * Inum (x # bs) s) / (2 * real_of_int n * real_of_int m)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2181
      \<in> (\<lambda>((t, n), s, m). (Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m) / 2) `
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2182
         (set U \<times> set U)"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2183
    using mnz nnz th
82292
5d91cca0aaf3 Tidied up a few messy proofs
paulson <lp15@cam.ac.uk>
parents: 74397
diff changeset
  2184
    by (force simp add: th add_divide_distrib algebra_simps split_def image_def)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2185
next
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2186
  fix t n s m
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2187
  assume tnU: "(t, n) \<in> set U" and smU: "(s, m) \<in> set U"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2188
  let ?N = "\<lambda>t. Inum (x # bs) t"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2189
  let ?st = "Add (Mul m t) (Mul n s)"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2190
  from Ul smU have mnz: "m \<noteq> 0"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2191
    by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2192
  from Ul tnU have nnz: "n \<noteq> 0"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2193
    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
  2194
  have st: "(?N t / real_of_int n + ?N s / real_of_int m) / 2 = ?N ?st / real_of_int (2 * n * m)"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2195
    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
  2196
  let ?P = "\<lambda>(t',n') (s',m'). (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
  2197
    (Inum (x # bs) t' / real_of_int n' + Inum (x # bs) s' / real_of_int m') / 2"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2198
  have Pc:"\<forall>a b. ?P a b = ?P b a"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2199
    by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2200
  from Ul alluopairs_set1 have Up:"\<forall>((t,n),(s,m)) \<in> set (alluopairs U). n \<noteq> 0 \<and> m \<noteq> 0"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2201
    by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2202
  from alluopairs_ex[OF Pc, where xs="U"] tnU smU
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2203
  have th':"\<exists>((t',n'),(s',m')) \<in> set (alluopairs U). ?P (t',n') (s',m')"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2204
    by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2205
  then obtain t' n' s' m' where ts'_U: "((t',n'),(s',m')) \<in> set (alluopairs U)"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2206
    and Pts': "?P (t', n') (s', m')"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2207
    by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2208
  from ts'_U Up have mnz': "m' \<noteq> 0" and nnz': "n'\<noteq> 0"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2209
    by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2210
  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
  2211
  have st': "(?N t' / real_of_int n' + ?N s' / real_of_int m') / 2 = ?N ?st' / real_of_int (2 * n' * m')"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2212
    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
  2213
  from Pts' have "(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
  2214
    (Inum (x # bs) t' / real_of_int n' + Inum (x # bs) s' / real_of_int m') / 2"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2215
    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
  2216
  also have "\<dots> = (\<lambda>(t, n). Inum (x # bs) t / real_of_int n)
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2217
      ((\<lambda>((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ((t', n'), (s', m')))"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2218
    by (simp add: st')
61609
77b453bd616f Coercion "real" now has type nat => real only 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
  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
  2220
    \<in> (\<lambda>(t, n). Inum (x # bs) t / real_of_int n) `
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2221
      (\<lambda>((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ` set (alluopairs U)"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2222
    using ts'_U by blast
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2223
qed
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2224
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2225
lemma uset_cong:
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2226
  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
  2227
    and UU': "((\<lambda>(t,n). Inum (x # bs) t / real_of_int n) ` U') =
77b453bd616f Coercion "real" now has type nat => real only 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
      ((\<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))"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2229
      (is "?f ` U' = ?g ` (U \<times> U)")
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2230
    and U: "\<forall>(t,n) \<in> U. numbound0 t \<and> n > 0"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2231
    and U': "\<forall>(t,n) \<in> U'. numbound0 t \<and> n > 0"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2232
  shows "(\<exists>(t,n) \<in> U. \<exists>(s,m) \<in> U. Ifm (x # bs) (usubst p (Add (Mul m t) (Mul n s), 2 * n * m))) =
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2233
    (\<exists>(t,n) \<in> U'. Ifm (x # bs) (usubst p (t, n)))"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2234
    (is "?lhs \<longleftrightarrow> ?rhs")
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2235
proof
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2236
  show ?rhs if ?lhs
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2237
  proof -
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2238
    from that obtain t n s m where tnU: "(t, n) \<in> U" and smU: "(s, m) \<in> U"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2239
      and Pst: "Ifm (x # bs) (usubst p (Add (Mul m t) (Mul n s), 2 * n * m))"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2240
      by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2241
    let ?N = "\<lambda>t. Inum (x#bs) t"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2242
    from tnU smU U have tnb: "numbound0 t" and np: "n > 0"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2243
      and snb: "numbound0 s" and mp: "m > 0"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2244
      by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2245
    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
  2246
    from np mp have mnp: "real_of_int (2 * n * m) > 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
  2247
      by (simp add: mult.commute of_int_mult[symmetric] del: of_int_mult)
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2248
    from tnb snb have stnb: "numbound0 ?st"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2249
      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
  2250
    have st: "(?N t / real_of_int n + ?N s / real_of_int m) / 2 = ?N ?st / real_of_int (2 * n * m)"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2251
      using mp np by (simp add: algebra_simps add_divide_distrib)
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2252
    from tnU smU UU' have "?g ((t, n), (s, m)) \<in> ?f ` U'"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2253
      by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2254
    then have "\<exists>(t',n') \<in> U'. ?g ((t, n), (s, m)) = ?f (t', n')"
82292
5d91cca0aaf3 Tidied up a few messy proofs
paulson <lp15@cam.ac.uk>
parents: 74397
diff changeset
  2255
      by fastforce
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2256
    then obtain t' n' where tnU': "(t',n') \<in> U'" and th: "?g ((t, n), (s, m)) = ?f (t', n')"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2257
      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
  2258
    from U' tnU' have tnb': "numbound0 t'" and np': "real_of_int n' > 0"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2259
      by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2260
    from usubst_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
  2261
    have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real_of_int (2 * n * m) # bs) p"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2262
      by simp
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2263
    from conjunct1[OF usubst_I[OF lp np' tnb', where bs="bs" and x="x"], symmetric]
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2264
      th[simplified split_def fst_conv snd_conv,symmetric] Pst2[simplified st[symmetric]]
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2265
    have "Ifm (x # bs) (usubst p (t', n'))"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2266
      by (simp only: st)
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2267
    then show ?thesis
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2268
      using tnU' by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2269
  qed
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2270
  show ?lhs if ?rhs
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2271
  proof -
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2272
    from that obtain t' n' where tnU': "(t', n') \<in> U'" and Pt': "Ifm (x # bs) (usubst p (t', n'))"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2273
      by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2274
    from tnU' UU' have "?f (t', n') \<in> ?g ` (U \<times> U)"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2275
      by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2276
    then have "\<exists>((t,n),(s,m)) \<in> U \<times> U. ?f (t', n') = ?g ((t, n), (s, m))"
82292
5d91cca0aaf3 Tidied up a few messy proofs
paulson <lp15@cam.ac.uk>
parents: 74397
diff changeset
  2277
      by force
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2278
    then obtain t n s m where tnU: "(t, n) \<in> U" and smU: "(s, m) \<in> U" and
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2279
      th: "?f (t', n') = ?g ((t, n), (s, m))"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2280
      by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2281
    let ?N = "\<lambda>t. Inum (x # bs) t"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2282
    from tnU smU U have tnb: "numbound0 t" and np: "n > 0"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2283
      and snb: "numbound0 s" and mp: "m > 0"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2284
      by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2285
    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
  2286
    from np mp have mnp: "real_of_int (2 * n * m) > 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
  2287
      by (simp add: mult.commute of_int_mult[symmetric] del: of_int_mult)
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2288
    from tnb snb have stnb: "numbound0 ?st"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2289
      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
  2290
    have st: "(?N t / real_of_int n + ?N s / real_of_int m) / 2 = ?N ?st / real_of_int (2 * n * m)"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2291
      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
  2292
    from U' tnU' have tnb': "numbound0 t'" and np': "real_of_int n' > 0"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2293
      by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2294
    from usubst_I[OF lp np' tnb', where bs="bs" and x="x",simplified
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2295
      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
  2296
    have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real_of_int (2 * n * m) # bs) p"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2297
      by simp
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2298
    with usubst_I[OF lp mnp stnb, where x="x" and bs="bs"] tnU smU
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2299
    show ?thesis by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2300
  qed
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2301
qed
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2302
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
  2303
lemma ferrack:
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2304
  assumes qf: "qfree p"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2305
  shows "qfree (ferrack p) \<and> (Ifm bs (ferrack p) \<longleftrightarrow> (\<exists>x. Ifm (x # bs) p))"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2306
  (is "_ \<and> (?rhs \<longleftrightarrow> ?lhs)")
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  2307
proof -
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2308
  let ?I = "\<lambda>x p. Ifm (x # bs) p"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2309
  fix x
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2310
  let ?N = "\<lambda>t. Inum (x # bs) t"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  2311
  let ?q = "rlfm (simpfm p)"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2312
  let ?U = "uset ?q"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2313
  let ?Up = "alluopairs ?U"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2314
  let ?g = "\<lambda>((t,n),(s,m)). (Add (Mul m t) (Mul n s), 2 * n * m)"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2315
  let ?S = "map ?g ?Up"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2316
  let ?SS = "map simp_num_pair ?S"
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  2317
  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
  2318
  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
  2319
  let ?h = "\<lambda>((t,n),(s,m)). (?N t / real_of_int n + ?N s / real_of_int m) / 2"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2320
  let ?F = "\<lambda>p. \<exists>a \<in> set (uset p). \<exists>b \<in> set (uset p). ?I x (usubst p (?g (a, b)))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  2321
  let ?ep = "evaldjf (simpfm \<circ> (usubst ?q)) ?Y"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2322
  from rlfm_I[OF simpfm_qf[OF qf]] have lq: "isrlfm ?q"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2323
    by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2324
  from alluopairs_set1[where xs="?U"] have UpU: "set ?Up \<subseteq> set ?U \<times> set ?U"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2325
    by simp
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  2326
  from uset_l[OF lq] have U_l: "\<forall>(t,n) \<in> set ?U. numbound0 t \<and> n > 0" .
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  2327
  from U_l UpU
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2328
  have "\<forall>((t,n),(s,m)) \<in> set ?Up. numbound0 t \<and> n> 0 \<and> numbound0 s \<and> m > 0"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2329
    by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2330
  then have Snb: "\<forall>(t,n) \<in> set ?S. numbound0 t \<and> n > 0 "
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2331
    by auto
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  2332
  have Y_l: "\<forall>(t,n) \<in> set ?Y. numbound0 t \<and> n > 0"
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  2333
  proof -
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2334
    have "numbound0 t \<and> n > 0" if tnY: "(t, n) \<in> set ?Y" for t n
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2335
    proof -
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2336
      from that have "(t,n) \<in> set ?SS"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2337
        by simp
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2338
      then have "\<exists>(t',n') \<in> set ?S. simp_num_pair (t', n') = (t, n)"
82292
5d91cca0aaf3 Tidied up a few messy proofs
paulson <lp15@cam.ac.uk>
parents: 74397
diff changeset
  2339
        by (force simp add: split_def simp del: map_map)
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2340
      then obtain t' n' where tn'S: "(t', n') \<in> set ?S" and tns: "simp_num_pair (t', n') = (t, n)"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2341
        by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2342
      from tn'S Snb have tnb: "numbound0 t'" and np: "n' > 0"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2343
        by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2344
      from simp_num_pair_l[OF tnb np tns] show ?thesis .
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2345
    qed
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  2346
    then show ?thesis by blast
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2347
  qed
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2348
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2349
  have YU: "(?f ` set ?Y) = (?h ` (set ?U \<times> set ?U))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  2350
  proof -
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2351
    from simp_num_pair_ci[where bs="x#bs"] have "\<forall>x. (?f \<circ> simp_num_pair) x = ?f x"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2352
      by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2353
    then have th: "?f \<circ> simp_num_pair = ?f"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2354
      by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2355
    have "(?f ` set ?Y) = ((?f \<circ> simp_num_pair) ` set ?S)"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2356
      by (simp add: comp_assoc image_comp)
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2357
    also have "\<dots> = ?f ` set ?S"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2358
      by (simp add: th)
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2359
    also have "\<dots> = (?f \<circ> ?g) ` set ?Up"
56154
f0a927235162 more complete set of lemmas wrt. image and composition
haftmann
parents: 55422
diff changeset
  2360
      by (simp only: set_map o_def image_comp)
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2361
    also have "\<dots> = ?h ` (set ?U \<times> set ?U)"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2362
      using uset_cong_aux[OF U_l, where x="x" and bs="bs", simplified set_map image_comp]
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2363
      by blast
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2364
    finally show ?thesis .
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2365
  qed
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2366
  have "\<forall>(t,n) \<in> set ?Y. bound0 (simpfm (usubst ?q (t, n)))"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  2367
  proof -
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2368
    have "bound0 (simpfm (usubst ?q (t, n)))" if tnY: "(t,n) \<in> set ?Y" for t n
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2369
    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
  2370
      from Y_l that have tnb: "numbound0 t" and np: "real_of_int n > 0"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2371
        by auto
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2372
      from usubst_I[OF lq np tnb] have "bound0 (usubst ?q (t, n))"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2373
        by simp
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2374
      then show ?thesis
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2375
        using simpfm_bound0 by simp
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2376
    qed
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  2377
    then show ?thesis by blast
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2378
  qed
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2379
  then have ep_nb: "bound0 ?ep"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2380
    using evaldjf_bound0[where xs="?Y" and f="simpfm \<circ> (usubst ?q)"] by auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2381
  let ?mp = "minusinf ?q"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2382
  let ?pp = "plusinf ?q"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2383
  let ?M = "?I x ?mp"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2384
  let ?P = "?I x ?pp"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2385
  let ?res = "disj ?mp (disj ?pp ?ep)"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2386
  from rminusinf_bound0[OF lq] rplusinf_bound0[OF lq] ep_nb have nbth: "bound0 ?res"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2387
    by auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2388
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2389
  from conjunct1[OF rlfm_I[OF simpfm_qf[OF qf]]] simpfm have th: "?lhs = (\<exists>x. ?I x ?q)"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2390
    by auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2391
  from th fr_equsubst[OF lq, where bs="bs" and x="x"] have lhfr: "?lhs = (?M \<or> ?P \<or> ?F ?q)"
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2392
    by (simp only: split_def fst_conv snd_conv)
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  2393
  also have "\<dots> = (?M \<or> ?P \<or> (\<exists>(t,n) \<in> set ?Y. ?I x (simpfm (usubst ?q (t,n)))))"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2394
    using uset_cong[OF lq YU U_l Y_l] by (simp only: split_def fst_conv snd_conv simpfm)
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2395
  also have "\<dots> = (Ifm (x#bs) ?res)"
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  2396
    using evaldjf_ex[where ps="?Y" and bs = "x#bs" and f="simpfm \<circ> (usubst ?q)",symmetric]
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60767
diff changeset
  2397
    by (simp add: split_def prod.collapse)
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2398
  finally have lheq: "?lhs = Ifm bs (decr ?res)"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2399
    using decr[OF nbth] by blast
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2400
  then have lr: "?lhs = ?rhs"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2401
    unfolding ferrack_def Let_def
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2402
    by (cases "?mp = T \<or> ?pp = T", auto) (simp add: disj_def)+
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2403
  from decr_qf[OF nbth] have "qfree (ferrack p)"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2404
    by (auto simp add: Let_def ferrack_def)
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2405
  with lr show ?thesis
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2406
    by blast
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2407
qed
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2408
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2409
definition linrqe:: "fm \<Rightarrow> fm"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2410
  where "linrqe p = qelim (prep p) ferrack"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2411
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2412
theorem linrqe: "Ifm bs (linrqe p) = Ifm bs p \<and> qfree (linrqe p)"
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2413
  using ferrack qelim_ci prep
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2414
  unfolding linrqe_def by auto
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2415
60711
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2416
definition ferrack_test :: "unit \<Rightarrow> fm"
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2417
where
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2418
  "ferrack_test u =
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2419
    linrqe (A (A (Imp (Lt (Sub (Bound 1) (Bound 0)))
799044496769 tuned proofs;
wenzelm
parents: 60710
diff changeset
  2420
      (E (Eq (Sub (Add (Bound 0) (Bound 2)) (Bound 1)))))))"
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2421
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 59621
diff changeset
  2422
ML_val \<open>@{code ferrack_test} ()\<close>
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2423
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 59621
diff changeset
  2424
oracle linr_oracle = \<open>
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2425
let
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2426
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
  2427
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: 49962
diff changeset
  2428
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: 49962
diff changeset
  2429
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
  2430
fun num_of_term vs (Free vT) = mk_Bound (find_index (fn vT' => vT = vT') vs)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69266
diff changeset
  2431
  | num_of_term vs \<^term>\<open>real_of_int (0::int)\<close> = mk_C 0
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69266
diff changeset
  2432
  | num_of_term vs \<^term>\<open>real_of_int (1::int)\<close> = mk_C 1
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2433
  | num_of_term vs \<^Const_>\<open>zero_class.zero \<^Type>\<open>real\<close>\<close> = mk_C 0
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2434
  | num_of_term vs \<^Const_>\<open>one_class.one \<^Type>\<open>real\<close>\<close> = mk_C 1
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
  2435
  | num_of_term vs (Bound i) = mk_Bound i
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2436
  | num_of_term vs \<^Const_>\<open>uminus \<^Type>\<open>real\<close> for t'\<close> = @{code Neg} (num_of_term vs t')
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2437
  | num_of_term vs \<^Const_>\<open>plus \<^Type>\<open>real\<close> for t1 t2\<close> =
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  2438
     @{code Add} (num_of_term vs t1, num_of_term vs t2)
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2439
  | num_of_term vs \<^Const_>\<open>minus \<^Type>\<open>real\<close> for t1 t2\<close> =
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  2440
     @{code Sub} (num_of_term vs t1, num_of_term vs t2)
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2441
  | num_of_term vs \<^Const_>\<open>times \<^Type>\<open>real\<close> for t1 t2\<close> = (case num_of_term vs t1
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2442
     of @{code C} i => @{code Mul} (i, num_of_term vs t2)
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  2443
      | _ => error "num_of_term: unsupported multiplication")
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69266
diff changeset
  2444
  | num_of_term vs (\<^term>\<open>real_of_int :: int \<Rightarrow> real\<close> $ t') =
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
  2445
     (mk_C (snd (HOLogic.dest_number t'))
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  2446
       handle TERM _ => error ("num_of_term: unknown term"))
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  2447
  | num_of_term vs t' =
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
  2448
     (mk_C (snd (HOLogic.dest_number t'))
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  2449
       handle TERM _ => error ("num_of_term: unknown term"));
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2450
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2451
fun fm_of_term vs \<^Const_>\<open>True\<close> = @{code T}
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2452
  | fm_of_term vs \<^Const_>\<open>False\<close> = @{code F}
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2453
  | fm_of_term vs \<^Const_>\<open>less \<^Type>\<open>real\<close> for t1 t2\<close> =
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  2454
      @{code Lt} (@{code Sub} (num_of_term vs t1, num_of_term vs t2))
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2455
  | fm_of_term vs \<^Const_>\<open>less_eq \<^Type>\<open>real\<close> for t1 t2\<close> =
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  2456
      @{code Le} (@{code Sub} (num_of_term vs t1, num_of_term vs t2))
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2457
  | fm_of_term vs \<^Const_>\<open>HOL.eq \<^Type>\<open>real\<close> for t1 t2\<close> =
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  2458
      @{code Eq} (@{code Sub} (num_of_term vs t1, num_of_term vs t2))
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2459
  | fm_of_term vs \<^Const_>\<open>HOL.eq \<^Type>\<open>bool\<close> for t1 t2\<close> =
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  2460
      @{code Iff} (fm_of_term vs t1, fm_of_term vs t2)
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2461
  | fm_of_term vs \<^Const_>\<open>HOL.conj for t1 t2\<close> = @{code And} (fm_of_term vs t1, fm_of_term vs t2)
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2462
  | fm_of_term vs \<^Const_>\<open>HOL.disj for t1 t2\<close> = @{code Or} (fm_of_term vs t1, fm_of_term vs t2)
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2463
  | fm_of_term vs \<^Const_>\<open>HOL.implies for t1 t2\<close> = @{code Imp} (fm_of_term vs t1, fm_of_term vs t2)
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2464
  | fm_of_term vs \<^Const_>\<open>HOL.Not for t'\<close> = @{code Not} (fm_of_term vs t')
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2465
  | fm_of_term vs \<^Const_>\<open>Ex _ for \<open>Abs (xn, xT, p)\<close>\<close> = @{code E} (fm_of_term (("", dummyT) :: vs) p)
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2466
  | fm_of_term vs \<^Const_>\<open>All _ for \<open>Abs (xn, xT, p)\<close>\<close> = @{code A} (fm_of_term (("", dummyT) ::  vs) p)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69266
diff changeset
  2467
  | fm_of_term vs t = error ("fm_of_term : unknown term " ^ Syntax.string_of_term \<^context> t);
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2468
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69266
diff changeset
  2469
fun term_of_num vs (@{code C} i) = \<^term>\<open>real_of_int :: int \<Rightarrow> real\<close> $
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
  2470
      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: 49962
diff changeset
  2471
  | term_of_num vs (@{code Bound} n) = Free (nth vs (@{code integer_of_nat} n))
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2472
  | term_of_num vs (@{code Neg} t') = \<^Const>\<open>uminus \<^Type>\<open>real\<close> for \<open>term_of_num vs t'\<close>\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2473
  | term_of_num vs (@{code Add} (t1, t2)) =
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2474
      \<^Const>\<open>plus \<^Type>\<open>real\<close> for \<open>term_of_num vs t1\<close> \<open>term_of_num vs t2\<close>\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2475
  | term_of_num vs (@{code Sub} (t1, t2)) =
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2476
      \<^Const>\<open>minus \<^Type>\<open>real\<close> for \<open>term_of_num vs t1\<close> \<open>term_of_num vs t2\<close>\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2477
  | term_of_num vs (@{code Mul} (i, t2)) =
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2478
      \<^Const>\<open>times \<^Type>\<open>real\<close> for \<open>term_of_num vs (@{code C} i)\<close> \<open>term_of_num vs t2\<close>\<close>
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2479
  | term_of_num vs (@{code CN} (n, i, t)) = term_of_num vs (@{code Add} (@{code Mul} (i, @{code Bound} n), t));
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2480
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2481
fun term_of_fm vs @{code T} = \<^Const>\<open>True\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2482
  | term_of_fm vs @{code F} = \<^Const>\<open>False\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2483
  | term_of_fm vs (@{code Lt} t) = \<^Const>\<open>less \<^Type>\<open>real\<close> for \<open>term_of_num vs t\<close> \<^term>\<open>0::real\<close>\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2484
  | term_of_fm vs (@{code Le} t) = \<^Const>\<open>less_eq \<^Type>\<open>real\<close> for \<open>term_of_num vs t\<close> \<^term>\<open>0::real\<close>\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2485
  | term_of_fm vs (@{code Gt} t) = \<^Const>\<open>less \<^Type>\<open>real\<close> for \<^term>\<open>0::real\<close> \<open>term_of_num vs t\<close>\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2486
  | term_of_fm vs (@{code Ge} t) = \<^Const>\<open>less_eq \<^Type>\<open>real\<close> for \<^term>\<open>0::real\<close> \<open>term_of_num vs t\<close>\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2487
  | term_of_fm vs (@{code Eq} t) = \<^Const>\<open>HOL.eq \<^Type>\<open>real\<close> for \<open>term_of_num vs t\<close> \<^term>\<open>0::real\<close>\<close>
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 73869
diff changeset
  2488
  | term_of_fm vs (@{code NEq} t) = term_of_fm vs (@{code Not} (@{code Eq} t))
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2489
  | term_of_fm vs (@{code Not} t') = \<^Const>\<open>HOL.Not for \<open>term_of_fm vs t'\<close>\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2490
  | term_of_fm vs (@{code And} (t1, t2)) = \<^Const>\<open>HOL.conj for \<open>term_of_fm vs t1\<close> \<open>term_of_fm vs t2\<close>\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2491
  | term_of_fm vs (@{code Or} (t1, t2)) = \<^Const>\<open>HOL.disj for \<open>term_of_fm vs t1\<close> \<open>term_of_fm vs t2\<close>\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2492
  | term_of_fm vs (@{code Imp} (t1, t2)) = \<^Const>\<open>HOL.implies for \<open>term_of_fm vs t1\<close> \<open>term_of_fm vs t2\<close>\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 74101
diff changeset
  2493
  | term_of_fm vs (@{code Iff} (t1, t2)) = \<^Const>\<open>HOL.eq \<^Type>\<open>bool\<close> for \<open>term_of_fm vs t1\<close> \<open>term_of_fm vs t2\<close>\<close>
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2494
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  2495
in fn (ctxt, t) =>
60710
07089a750d2a tuned proofs;
wenzelm
parents: 60533
diff changeset
  2496
  let
36853
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  2497
    val vs = Term.add_frees t [];
c8e4102b08aa modernized specifications; tuned reification
haftmann
parents: 35416
diff changeset
  2498
    val t' = (term_of_fm vs o @{code linrqe} o fm_of_term vs) t;
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59580
diff changeset
  2499
  in (Thm.cterm_of ctxt o HOLogic.mk_Trueprop o HOLogic.mk_eq) (t, t') end
69266
7cc2d66a92a6 proper ML expressions, without trailing semicolons;
wenzelm
parents: 69064
diff changeset
  2500
end
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 59621
diff changeset
  2501
\<close>
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2502
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69597
diff changeset
  2503
ML_file \<open>ferrack_tac.ML\<close>
47432
e1576d13e933 more standard method setup;
wenzelm
parents: 47142
diff changeset
  2504
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 59621
diff changeset
  2505
method_setup rferrack = \<open>
53168
d998de7f0efc tuned signature;
wenzelm
parents: 51272
diff changeset
  2506
  Scan.lift (Args.mode "no_quantify") >>
47432
e1576d13e933 more standard method setup;
wenzelm
parents: 47142
diff changeset
  2507
    (fn q => fn ctxt => SIMPLE_METHOD' (Ferrack_Tac.linr_tac ctxt (not q)))
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 59621
diff changeset
  2508
\<close> "decision procedure for linear real arithmetic"
47432
e1576d13e933 more standard method setup;
wenzelm
parents: 47142
diff changeset
  2509
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2510
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2511
lemma
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2512
  fixes x :: real
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2513
  shows "2 * x \<le> 2 * x \<and> 2 * x \<le> 2 * x + 1"
49070
f00fee6d21d4 tuned proofs;
wenzelm
parents: 48891
diff changeset
  2514
  by rferrack
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2515
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2516
lemma
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2517
  fixes x :: real
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2518
  shows "\<exists>y \<le> x. x = y + 1"
49070
f00fee6d21d4 tuned proofs;
wenzelm
parents: 48891
diff changeset
  2519
  by rferrack
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2520
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2521
lemma
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2522
  fixes x :: real
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2523
  shows "\<not> (\<exists>z. x + z = x + z + 1)"
49070
f00fee6d21d4 tuned proofs;
wenzelm
parents: 48891
diff changeset
  2524
  by rferrack
29789
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2525
b4534c3e68f6 established session HOL-Reflection
haftmann
parents:
diff changeset
  2526
end