src/HOL/Complex/ex/linreif.ML
author wenzelm
Thu, 27 Mar 2008 14:41:07 +0100
changeset 26423 8408edac8f6b
parent 25938 2c1c0e989615
child 26939 1035c89b4c02
permissions -rw-r--r--
removed Display.raw_string_of_XXX (use regular Sign.string_of_XXX);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23515
3e7f62e68fe4 new code generator framework
haftmann
parents: 23317
diff changeset
     1
(*  ID:         $Id$
3e7f62e68fe4 new code generator framework
haftmann
parents: 23317
diff changeset
     2
    Author:     Amine Chaieb, TU Muenchen
3e7f62e68fe4 new code generator framework
haftmann
parents: 23317
diff changeset
     3
3e7f62e68fe4 new code generator framework
haftmann
parents: 23317
diff changeset
     4
The oracle for Mixed Real-Integer auantifier elimination
3e7f62e68fe4 new code generator framework
haftmann
parents: 23317
diff changeset
     5
based on the verified Code in ~/work/MIR/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
     6
*)
324622260d29 Added 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
structure ReflectedFerrack =
324622260d29 Added 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
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
    10
324622260d29 Added 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
open Ferrack;
23808
4e4b92e76219 fixed SML/NJ int problem
haftmann
parents: 23515
diff changeset
    12
23264
324622260d29 Added 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
exception LINR;
324622260d29 Added 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
324622260d29 Added 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
(* pseudo reification : term -> intterm *)
324622260d29 Added 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
val rT = Type ("RealDef.real",[]);
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    17
val bT = HOLogic.boolT;
23881
851c74f1bb69 moved class ord from Orderings.thy to HOL.thy
haftmann
parents: 23808
diff changeset
    18
val realC = @{term "RealDef.real :: int => real"};
851c74f1bb69 moved class ord from Orderings.thy to HOL.thy
haftmann
parents: 23808
diff changeset
    19
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
    20
23881
851c74f1bb69 moved class ord from Orderings.thy to HOL.thy
haftmann
parents: 23808
diff changeset
    21
fun num_of_term vs t = case t
851c74f1bb69 moved class ord from Orderings.thy to HOL.thy
haftmann
parents: 23808
diff changeset
    22
   of Free(xn,xT) => (case AList.lookup (op =) vs t of 
851c74f1bb69 moved class ord from Orderings.thy to HOL.thy
haftmann
parents: 23808
diff changeset
    23
           NONE   => error "Variable not found in the list!"
851c74f1bb69 moved class ord from Orderings.thy to HOL.thy
haftmann
parents: 23808
diff changeset
    24
         | 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
    25
      | 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
    26
      | 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
    27
      | @{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
    28
      | @{term "0::real"} => C 1
25938
2c1c0e989615 Efficient_Nat streamlined and improved
haftmann
parents: 25919
diff changeset
    29
      | Term.Bound i => Bound i
23881
851c74f1bb69 moved class ord from Orderings.thy to HOL.thy
haftmann
parents: 23808
diff changeset
    30
      | Const (@{const_name "HOL.uminus"},_)$t' => Neg (num_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
    31
      | 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
    32
      | Const (@{const_name "HOL.minus"},_)$t1$t2 => Sub (num_of_term vs t1,num_of_term vs t2)
23881
851c74f1bb69 moved class ord from Orderings.thy to HOL.thy
haftmann
parents: 23808
diff changeset
    33
      | Const (@{const_name "HOL.times"},_)$t1$t2 => (case (num_of_term vs t1) of C i => 
851c74f1bb69 moved class ord from Orderings.thy to HOL.thy
haftmann
parents: 23808
diff changeset
    34
          Mul (i,num_of_term vs t2)
851c74f1bb69 moved class ord from Orderings.thy to HOL.thy
haftmann
parents: 23808
diff changeset
    35
        | _ => error "num_of_term: unsupported Multiplication")
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 24630
diff changeset
    36
      | Const("RealDef.real",_) $ Const (@{const_name "Int.number_of"},_)$t' => C (HOLogic.dest_numeral t')
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 24630
diff changeset
    37
      | Const (@{const_name "Int.number_of"},_)$t' => C (HOLogic.dest_numeral t')
26423
8408edac8f6b removed Display.raw_string_of_XXX (use regular Sign.string_of_XXX);
wenzelm
parents: 25938
diff changeset
    38
      | _ => error ("num_of_term: unknown term " ^ Sign.string_of_term CPure.thy 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
    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 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    43
	Const("True",_) => 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
    44
      | Const("False",_) => F
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 23881
diff 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: 23881
diff changeset
    46
      | Const(@{const_name HOL.less_eq},_)$t1$t2 => Le (Sub (num_of_term vs 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
    47
      | Const("op =",eqT)$t1$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
    48
	if (domain_type eqT = rT)
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 23881
diff changeset
    49
	then Eq (Sub (num_of_term vs t1,num_of_term vs t2)) 
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 23881
diff changeset
    50
	else Iff(fm_of_term vs t1,fm_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
    51
      | Const("op &",_)$t1$t2 => And(fm_of_term vs t1,fm_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
    52
      | 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: 23881
diff changeset
    53
      | Const("op -->",_)$t1$t2 => Imp(fm_of_term vs t1,fm_of_term vs t2)
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 23881
diff changeset
    54
      | 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
    55
      | Const("Ex",_)$Term.Abs(xn,xT,p) => 
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: 23264
diff changeset
    56
	E(fm_of_term (map (fn(v,n) => (v,1+ 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
    57
      | Const("All",_)$Term.Abs(xn,xT,p) => 
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: 23264
diff changeset
    58
	A(fm_of_term (map (fn(v,n) => (v,1+ n)) vs) p)
26423
8408edac8f6b removed Display.raw_string_of_XXX (use regular Sign.string_of_XXX);
wenzelm
parents: 25938
diff changeset
    59
      | _ => error ("fm_of_term : unknown term!" ^ Sign.string_of_term CPure.thy 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
    60
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    61
324622260d29 Added 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
fun start_vs t =
25938
2c1c0e989615 Efficient_Nat streamlined and improved
haftmann
parents: 25919
diff changeset
    63
  let
2c1c0e989615 Efficient_Nat streamlined and improved
haftmann
parents: 25919
diff changeset
    64
    val fs = term_frees t
2c1c0e989615 Efficient_Nat streamlined and improved
haftmann
parents: 25919
diff changeset
    65
  in fs ~~ (0 upto (length fs - 1)) end;
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    66
324622260d29 Added 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
(* 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
    68
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    69
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
    70
    case l of
324622260d29 Added 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
	[] => NONE
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    72
      | (x,v')::xs => if v = v' then SOME x
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    73
		      else myassoc2 xs 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
    74
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    75
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
    76
    case t of 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    77
	C i => realC $ (HOLogic.mk_number HOLogic.intT i)
23881
851c74f1bb69 moved class ord from Orderings.thy to HOL.thy
haftmann
parents: 23808
diff changeset
    78
      | Bound n => the (myassoc2 vs n)
851c74f1bb69 moved class ord from Orderings.thy to HOL.thy
haftmann
parents: 23808
diff changeset
    79
      | Neg t' => Const(@{const_name HOL.uminus},rT --> rT)$(term_of_num vs t')
851c74f1bb69 moved class ord from Orderings.thy to HOL.thy
haftmann
parents: 23808
diff changeset
    80
      | Add(t1,t2) => Const(@{const_name HOL.plus},[rT,rT] ---> rT)$
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    81
			   (term_of_num vs t1)$(term_of_num vs t2)
23881
851c74f1bb69 moved class ord from Orderings.thy to HOL.thy
haftmann
parents: 23808
diff changeset
    82
      | Sub(t1,t2) => Const(@{const_name HOL.minus},[rT,rT] ---> rT)$
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    83
			   (term_of_num vs t1)$(term_of_num vs t2)
23881
851c74f1bb69 moved class ord from Orderings.thy to HOL.thy
haftmann
parents: 23808
diff changeset
    84
      | Mul(i,t2) => Const(@{const_name HOL.times},[rT,rT] ---> rT)$
23264
324622260d29 Added 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
			   (term_of_num vs (C i))$(term_of_num vs t2)
23515
3e7f62e68fe4 new code generator framework
haftmann
parents: 23317
diff changeset
    86
      | Cn(n,i,t) => term_of_num vs (Add (Mul(i,Bound n),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
    87
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    88
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
    89
    case t of 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    90
	T => HOLogic.true_const 
324622260d29 Added 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
      | F => HOLogic.false_const
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 23881
diff changeset
    92
      | Lt t => Const(@{const_name HOL.less},[rT,rT] ---> bT)$
23264
324622260d29 Added 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
			   (term_of_num vs t)$ rzero
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 23881
diff changeset
    94
      | Le t => Const(@{const_name HOL.less_eq},[rT,rT] ---> bT)$
23264
324622260d29 Added 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
			  (term_of_num vs t)$ rzero
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 23881
diff changeset
    96
      | Gt t => Const(@{const_name HOL.less},[rT,rT] ---> bT)$
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    97
			   rzero $(term_of_num vs t)
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 23881
diff changeset
    98
      | Ge t => Const(@{const_name HOL.less_eq},[rT,rT] ---> bT)$
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    99
			  rzero $(term_of_num vs t)
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 23881
diff changeset
   100
      | Eq t => Const("op =",[rT,rT] ---> bT)$
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   101
			   (term_of_num vs t)$ rzero
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 23881
diff changeset
   102
      | NEq t => term_of_fm vs (Not (Eq t))
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 23881
diff changeset
   103
      | 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
   104
      | 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
   105
      | 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: 23881
diff changeset
   106
      | Imp(t1,t2) => HOLogic.imp$(term_of_fm vs t1)$(term_of_fm vs t2)
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 23881
diff changeset
   107
      | Iff(t1,t2) => (HOLogic.eq_const bT)$(term_of_fm vs t1)$
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   108
					   (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
   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 linrqe_oracle thy t = 
25938
2c1c0e989615 Efficient_Nat streamlined and improved
haftmann
parents: 25919
diff changeset
   114
  let 
2c1c0e989615 Efficient_Nat streamlined and improved
haftmann
parents: 25919
diff changeset
   115
    val vs = start_vs t
2c1c0e989615 Efficient_Nat streamlined and improved
haftmann
parents: 25919
diff changeset
   116
  in HOLogic.mk_Trueprop (HOLogic.mk_eq (t, term_of_fm vs (linrqe (fm_of_term vs t)))) end;
23515
3e7f62e68fe4 new code generator framework
haftmann
parents: 23317
diff changeset
   117
23264
324622260d29 Added 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
end;