| author | haftmann | 
| Thu, 18 Oct 2007 09:20:59 +0200 | |
| changeset 25079 | db5fdc34f3af | 
| parent 24423 | ae9cd0e92423 | 
| child 25919 | 8b1c0d434824 | 
| permissions | -rw-r--r-- | 
| 23858 | 1 | (* Title: HOL/Complex/ex/mireif.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Amine Chaieb, TU Muenchen | |
| 4 | ||
| 5 | Oracle for Mixed Real-Integer auantifier elimination | |
| 6 | based on the verified code in HOL/Complex/ex/MIR.thy. | |
| 23264 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 7 | *) | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 8 | |
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 9 | structure ReflectedMir = | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 10 | struct | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 11 | |
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 12 | open Mir; | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 13 | |
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 14 | exception MIR; | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 15 | |
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 16 | fun num_of_term vs t = | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 17 | case t of | 
| 23858 | 18 | Free(xn,xT) => (case AList.lookup (op =) vs t of | 
| 19 | NONE => error "Variable not found in the list!" | |
| 20 | | SOME n => Bound n) | |
| 23264 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 21 |       | Const("RealDef.real",_)$ @{term "0::int"} => C 0
 | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 22 |       | Const("RealDef.real",_)$ @{term "1::int"} => C 1
 | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 23 |       | @{term "0::real"} => C 0
 | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 24 |       | @{term "1::real"} => C 1
 | 
| 23317 
90be000da2a7
Temporarily use int instead of IntInf.int but Code generator should map HOL's int to ML's IntInf.int --- To be fixed
 chaieb parents: 
23264diff
changeset | 25 | | Term.Bound i => Bound (nat i) | 
| 23264 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 26 |       | Const(@{const_name "HOL.uminus"},_)$t' => Neg (num_of_term vs t')
 | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 27 |       | Const (@{const_name "HOL.plus"},_)$t1$t2 => Add (num_of_term vs t1,num_of_term vs t2)
 | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 28 |       | Const (@{const_name "HOL.minus"},_)$t1$t2 => Sub (num_of_term vs t1,num_of_term vs t2)
 | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 29 |       | Const (@{const_name "HOL.times"},_)$t1$t2 => 
 | 
| 23858 | 30 | (case (num_of_term vs t1) of C i => | 
| 31 | Mul (i,num_of_term vs t2) | |
| 32 | | _ => error "num_of_term: unsupported Multiplication") | |
| 23264 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 33 |       | Const("RealDef.real",_)$ (Const (@{const_name "RComplete.floor"},_)$ t') => Floor (num_of_term vs t')
 | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 34 |       | Const("RealDef.real",_)$ (Const (@{const_name "RComplete.ceiling"},_)$ t') => Neg(Floor (Neg (num_of_term vs t')))
 | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 35 |       | Const("RealDef.real",_) $ Const (@{const_name "Numeral.number_of"},_)$t' => C (HOLogic.dest_numeral t')
 | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 36 |       | Const (@{const_name "Numeral.number_of"},_)$t' => C (HOLogic.dest_numeral t')
 | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 37 |       | _ => error ("num_of_term: unknown term " ^ (Display.raw_string_of_term t));
 | 
| 23858 | 38 | |
| 23264 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 39 | |
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 40 | (* pseudo reification : term -> fm *) | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 41 | fun fm_of_term vs t = | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 42 | case t of | 
| 23858 | 43 |         Const("True",_) => T
 | 
| 23264 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 44 |       | Const("False",_) => F
 | 
| 24423 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
23881diff
changeset | 45 |       | Const(@{const_name HOL.less},_)$t1$t2 => Lt (Sub (num_of_term vs t1,num_of_term vs t2))
 | 
| 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
23881diff
changeset | 46 |       | Const(@{const_name HOL.less_eq},_)$t1$t2 => Le (Sub (num_of_term vs t1,num_of_term vs t2))
 | 
| 23858 | 47 |       | Const (@{const_name "MIR.rdvd"},_ )$ (Const("RealDef.real",_) $ (Const(@{const_name "Numeral.number_of"},_)$t1))$t2 => 
 | 
| 24423 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
23881diff
changeset | 48 | Dvd (HOLogic.dest_numeral t1, num_of_term vs t2) | 
| 23264 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 49 |       | Const("op =",eqT)$t1$t2 => 
 | 
| 23858 | 50 |         if (domain_type eqT = @{typ real})
 | 
| 24423 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
23881diff
changeset | 51 | then Eq (Sub (num_of_term vs t1, num_of_term vs t2)) | 
| 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
23881diff
changeset | 52 | else Iff (fm_of_term vs t1, fm_of_term vs t2) | 
| 23858 | 53 |       | Const("op &",_)$t1$t2 => And (fm_of_term vs t1, fm_of_term vs t2)
 | 
| 54 |       | Const("op |",_)$t1$t2 => Or (fm_of_term vs t1, fm_of_term vs t2)
 | |
| 24423 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
23881diff
changeset | 55 |       | Const("op -->",_)$t1$t2 => Imp (fm_of_term vs t1, fm_of_term vs t2)
 | 
| 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
23881diff
changeset | 56 |       | Const("Not",_)$t' => Not (fm_of_term vs t')
 | 
| 23264 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 57 |       | Const("Ex",_)$Abs(xn,xT,p) => 
 | 
| 23858 | 58 | E (fm_of_term (map (fn (v, n) => (v, Suc n)) vs) p) | 
| 23264 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 59 |       | Const("All",_)$Abs(xn,xT,p) => 
 | 
| 23858 | 60 | A (fm_of_term (map (fn(v, n) => (v, Suc n)) vs) p) | 
| 61 |       | _ => error ("fm_of_term : unknown term!" ^ Display.raw_string_of_term t);
 | |
| 23264 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 62 | |
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 63 | fun start_vs t = | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 64 | let val fs = term_frees t | 
| 23858 | 65 | in fs ~~ map nat (0 upto (length fs - 1)) | 
| 23264 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 66 | end ; | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 67 | |
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 68 | (* transform num and fm back to terms *) | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 69 | |
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 70 | fun myassoc2 l v = | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 71 | case l of | 
| 23858 | 72 | [] => NONE | 
| 23264 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 73 | | (x,v')::xs => if v = v' then SOME x | 
| 23858 | 74 | else myassoc2 xs v; | 
| 75 | ||
| 76 | val realC = @{term "real :: int => _"};
 | |
| 77 | val rzero = @{term "0::real"};
 | |
| 23264 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 78 | |
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 79 | fun term_of_num vs t = | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 80 | case t of | 
| 23858 | 81 | C i => realC $ (HOLogic.mk_number HOLogic.intT i) | 
| 23264 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 82 | | Bound n => valOf (myassoc2 vs n) | 
| 23858 | 83 |       | Neg (Floor (Neg t')) => realC $ (@{term "ceiling"} $ term_of_num vs t')
 | 
| 23264 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 84 |       | Neg t' => @{term "uminus:: real => _"} $ term_of_num vs t'
 | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 85 |       | Add(t1,t2) => @{term "op +:: real => _"} $ term_of_num vs t1 $ term_of_num vs t2
 | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 86 |       | Sub(t1,t2) => @{term "op -:: real => _"} $ term_of_num vs t1 $ term_of_num vs t2
 | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 87 |       | Mul(i,t2) => @{term "op -:: real => _"} $ term_of_num vs (C i) $ term_of_num vs t2
 | 
| 23858 | 88 |       | Floor t => realC $ (@{term "floor"} $ term_of_num vs t)
 | 
| 89 | | Cn(n,i,t) => term_of_num vs (Add(Mul(i,Bound n),t)) | |
| 90 | | Cf(c,t,s) => term_of_num vs (Add(Mul(c,Floor t),s)); | |
| 23264 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 91 | |
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 92 | fun term_of_fm vs t = | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 93 | case t of | 
| 23858 | 94 | T => HOLogic.true_const | 
| 23264 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 95 | | F => HOLogic.false_const | 
| 24423 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
23881diff
changeset | 96 |       | Lt t => @{term "op <:: real => _"} $ term_of_num vs t $ rzero
 | 
| 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
23881diff
changeset | 97 |       | Le t => @{term "op <=:: real => _"} $ term_of_num vs t $ rzero
 | 
| 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
23881diff
changeset | 98 |       | Gt t => @{term "op <:: real => _"}$ rzero $ term_of_num vs t
 | 
| 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
23881diff
changeset | 99 |       | Ge t => @{term "op <=:: real => _"} $ rzero $ term_of_num vs t
 | 
| 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
23881diff
changeset | 100 |       | Eq t => @{term "op = :: real => _"}$ term_of_num vs t $ rzero
 | 
| 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
23881diff
changeset | 101 | | NEq t => term_of_fm vs (Not (Eq t)) | 
| 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
23881diff
changeset | 102 | | NDvd (i,t) => term_of_fm vs (Not (Dvd (i,t))) | 
| 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
23881diff
changeset | 103 |       | Dvd (i,t) => @{term "op rdvd"} $ term_of_num vs (C i) $ term_of_num vs t
 | 
| 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
23881diff
changeset | 104 | | Not t' => HOLogic.Not$(term_of_fm vs t') | 
| 23264 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 105 | | And(t1,t2) => HOLogic.conj $ term_of_fm vs t1 $ term_of_fm vs t2 | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 106 | | Or(t1,t2) => HOLogic.disj $ term_of_fm vs t1 $ term_of_fm vs t2 | 
| 24423 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
23881diff
changeset | 107 | | Imp(t1,t2) => HOLogic.imp $ term_of_fm vs t1 $ term_of_fm vs t2 | 
| 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
23881diff
changeset | 108 | | Iff(t1,t2) => HOLogic.mk_eq (term_of_fm vs t1, term_of_fm vs t2) | 
| 23264 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 109 | | _ => error "If this is raised, Isabelle/HOL or generate_code is inconsistent!"; | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 110 | |
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 111 | (* The oracle *) | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 112 | |
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 113 | fun mircfr_oracle thy t = | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 114 | let | 
| 23858 | 115 | val vs = start_vs t | 
| 23264 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 116 | in HOLogic.mk_Trueprop (HOLogic.mk_eq(t, term_of_fm vs (mircfrqe (fm_of_term vs t)))) | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 117 | end; | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 118 | |
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 119 | fun mirlfr_oracle thy t = | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 120 | let | 
| 23858 | 121 | val vs = start_vs t | 
| 23264 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 122 | in HOLogic.mk_Trueprop (HOLogic.mk_eq(t, term_of_fm vs (mirlfrqe (fm_of_term vs t)))) | 
| 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 123 | end; | 
| 23858 | 124 | |
| 23264 
324622260d29
Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
 chaieb parents: diff
changeset | 125 | end; |