src/HOL/Complex/ex/ReflectedFerrack.thy
changeset 27436 9581777503e9
parent 27368 9f90ac19e32b
child 27456 52c7c42e7e27
equal deleted inserted replaced
27435:b3f8e9bdf9a7 27436:9581777503e9
     4 
     4 
     5 header {* Quatifier elimination for R(0,1,+,<) *}
     5 header {* Quatifier elimination for R(0,1,+,<) *}
     6 
     6 
     7 theory ReflectedFerrack
     7 theory ReflectedFerrack
     8 imports Main GCD Real Efficient_Nat
     8 imports Main GCD Real Efficient_Nat
     9 uses ("linreif.ML") ("linrtac.ML")
     9 uses ("linrtac.ML")
    10 begin
    10 begin
    11 
    11 
    12 
    12 
    13   (*********************************************************************************)
    13   (*********************************************************************************)
    14   (*          SOME GENERAL STUFF< HAS TO BE MOVED IN SOME LIB                      *)
    14   (*          SOME GENERAL STUFF< HAS TO BE MOVED IN SOME LIB                      *)
  1983     by (cases "?mp = T \<or> ?pp = T", auto) (simp add: disj_def)+
  1983     by (cases "?mp = T \<or> ?pp = T", auto) (simp add: disj_def)+
  1984   from decr_qf[OF nbth] have "qfree (ferrack p)" by (auto simp add: Let_def ferrack_def)
  1984   from decr_qf[OF nbth] have "qfree (ferrack p)" by (auto simp add: Let_def ferrack_def)
  1985   with lr show ?thesis by blast
  1985   with lr show ?thesis by blast
  1986 qed
  1986 qed
  1987 
  1987 
  1988 constdefs linrqe:: "fm \<Rightarrow> fm"
  1988 definition linrqe:: "fm \<Rightarrow> fm" where
  1989   "linrqe \<equiv> (\<lambda> p. qelim (prep p) ferrack)"
  1989   "linrqe p = qelim (prep p) ferrack"
  1990 
  1990 
  1991 theorem linrqe: "(Ifm bs (linrqe p) = Ifm bs p) \<and> qfree (linrqe p)"
  1991 theorem linrqe: "Ifm bs (linrqe p) = Ifm bs p \<and> qfree (linrqe p)"
  1992 using ferrack qelim_ci prep
  1992 using ferrack qelim_ci prep
  1993 unfolding linrqe_def by auto
  1993 unfolding linrqe_def by auto
  1994 
  1994 
  1995 definition
  1995 definition ferrack_test :: "unit \<Rightarrow> fm" where
  1996   ferrack_test :: "unit \<Rightarrow> fm"
       
  1997 where
       
  1998   "ferrack_test u = linrqe (A (A (Imp (Lt (Sub (Bound 1) (Bound 0)))
  1996   "ferrack_test u = linrqe (A (A (Imp (Lt (Sub (Bound 1) (Bound 0)))
  1999     (E (Eq (Sub (Add (Bound 0) (Bound 2)) (Bound 1)))))))"
  1997     (E (Eq (Sub (Add (Bound 0) (Bound 2)) (Bound 1)))))))"
  2000 
  1998 
  2001 export_code linrqe ferrack_test in SML module_name Ferrack
  1999 ML {* @{code ferrack_test} () *}
  2002 
  2000 
  2003 (*code_module Ferrack
  2001 oracle linr_oracle ("term") = {*
  2004   contains
  2002 
  2005     linrqe = linrqe
  2003 let
  2006     test = ferrack_test*)
  2004 
  2007 
  2005 fun num_of_term vs (t as Free (xn, xT)) = (case AList.lookup (op =) vs t
  2008 ML {* Ferrack.ferrack_test () *}
  2006      of NONE => error "Variable not found in the list!"
  2009 
  2007       | SOME n => @{code Bound} n)
  2010 use "linreif.ML"
  2008   | num_of_term vs @{term "real (0::int)"} = @{code C} 0
  2011 oracle linr_oracle ("term") = ReflectedFerrack.linrqe_oracle
  2009   | num_of_term vs @{term "real (1::int)"} = @{code C} 1
       
  2010   | num_of_term vs @{term "0::real"} = @{code C} 0
       
  2011   | num_of_term vs @{term "1::real"} = @{code C} 1
       
  2012   | num_of_term vs (Term.Bound i) = @{code Bound} i
       
  2013   | num_of_term vs (@{term "uminus :: real \<Rightarrow> real"} $ t') = @{code Neg} (num_of_term vs t')
       
  2014   | num_of_term vs (@{term "op + :: real \<Rightarrow> real \<Rightarrow> real"} $ t1 $ t2) = @{code Add} (num_of_term vs t1, num_of_term vs t2)
       
  2015   | num_of_term vs (@{term "op - :: real \<Rightarrow> real \<Rightarrow> real"} $ t1 $ t2) = @{code Sub} (num_of_term vs t1, num_of_term vs t2)
       
  2016   | num_of_term vs (@{term "op * :: real \<Rightarrow> real \<Rightarrow> real"} $ t1 $ t2) = (case (num_of_term vs t1)
       
  2017      of @{code C} i => @{code Mul} (i, num_of_term vs t2)
       
  2018       | _ => error "num_of_term: unsupported Multiplication")
       
  2019   | num_of_term vs (@{term "real :: int \<Rightarrow> real"} $ (@{term "number_of :: int \<Rightarrow> int"} $ t')) = @{code C} (HOLogic.dest_numeral t')
       
  2020   | num_of_term vs (@{term "number_of :: int \<Rightarrow> int"} $ t') = @{code C} (HOLogic.dest_numeral t')
       
  2021   | num_of_term vs t = error ("num_of_term: unknown term " ^ Syntax.string_of_term_global @{theory} t);
       
  2022 
       
  2023 fun fm_of_term vs @{term True} = @{code T}
       
  2024   | fm_of_term vs @{term False} = @{code T}
       
  2025   | fm_of_term vs (@{term "op < :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) = @{code Lt} (@{code Sub} (num_of_term vs t1, num_of_term vs t2))
       
  2026   | fm_of_term vs (@{term "op \<le> :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) = @{code Le} (@{code Sub} (num_of_term vs t1, num_of_term vs t2))
       
  2027   | fm_of_term vs (@{term "op = :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) = @{code Eq} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) 
       
  2028   | fm_of_term vs (@{term "op = :: bool \<Rightarrow> bool \<Rightarrow> bool"} $ t1 $ t2) = @{code Iff} (fm_of_term vs t1, fm_of_term vs t2)
       
  2029   | fm_of_term vs (@{term "op &"} $ t1 $ t2) = @{code And} (fm_of_term vs t1, fm_of_term vs t2)
       
  2030   | fm_of_term vs (@{term "op |"} $ t1 $ t2) = @{code Or} (fm_of_term vs t1, fm_of_term vs t2)
       
  2031   | fm_of_term vs (@{term "op -->"} $ t1 $ t2) = @{code Imp} (fm_of_term vs t1, fm_of_term vs t2)
       
  2032   | fm_of_term vs (@{term "Not"} $ t') = @{code NOT} (fm_of_term vs t')
       
  2033   | fm_of_term vs (Const("Ex", _) $ Term.Abs (xn, xT, p)) =
       
  2034       @{code E} (fm_of_term (map (fn (v, n) => (v, n + 1)) vs) p)
       
  2035   | fm_of_term vs (Const("All", _) $ Term.Abs(xn, xT, p)) =
       
  2036       @{code A} (fm_of_term (map (fn (v, n) => (v, n + 1)) vs) p)
       
  2037   | fm_of_term vs t = error ("fm_of_term : unknown term " ^ Syntax.string_of_term_global @{theory} t);
       
  2038 
       
  2039 fun term_of_num vs (@{code C} i) = @{term "real :: int \<Rightarrow> real"} $ HOLogic.mk_number HOLogic.intT i
       
  2040   | term_of_num vs (@{code Bound} n) = fst (the (find_first (fn (_, m) => n = m) vs))
       
  2041   | term_of_num vs (@{code Neg} t') = @{term "uminus :: real \<Rightarrow> real"} $ term_of_num vs t'
       
  2042   | term_of_num vs (@{code Add} (t1, t2)) = @{term "op + :: real \<Rightarrow> real \<Rightarrow> real"} $
       
  2043       term_of_num vs t1 $ term_of_num vs t2
       
  2044   | term_of_num vs (@{code Sub} (t1, t2)) = @{term "op - :: real \<Rightarrow> real \<Rightarrow> real"} $
       
  2045       term_of_num vs t1 $ term_of_num vs t2
       
  2046   | term_of_num vs (@{code Mul} (i, t2)) = @{term "op * :: real \<Rightarrow> real \<Rightarrow> real"} $
       
  2047       term_of_num vs (@{code C} i) $ term_of_num vs t2
       
  2048   | term_of_num vs (@{code CN} (n, i, t)) = term_of_num vs (@{code Add} (@{code Mul} (i, @{code Bound} n), t));
       
  2049 
       
  2050 fun term_of_fm vs @{code T} = HOLogic.true_const 
       
  2051   | term_of_fm vs @{code F} = HOLogic.false_const
       
  2052   | term_of_fm vs (@{code Lt} t) = @{term "op < :: real \<Rightarrow> real \<Rightarrow> bool"} $
       
  2053       term_of_num vs t $ @{term "0::real"}
       
  2054   | term_of_fm vs (@{code Le} t) = @{term "op \<le> :: real \<Rightarrow> real \<Rightarrow> bool"} $
       
  2055       term_of_num vs t $ @{term "0::real"}
       
  2056   | term_of_fm vs (@{code Gt} t) = @{term "op < :: real \<Rightarrow> real \<Rightarrow> bool"} $
       
  2057       @{term "0::real"} $ term_of_num vs t
       
  2058   | term_of_fm vs (@{code Ge} t) = @{term "op \<le> :: real \<Rightarrow> real \<Rightarrow> bool"} $
       
  2059       @{term "0::real"} $ term_of_num vs t
       
  2060   | term_of_fm vs (@{code Eq} t) = @{term "op = :: real \<Rightarrow> real \<Rightarrow> bool"} $
       
  2061       term_of_num vs t $ @{term "0::real"}
       
  2062   | term_of_fm vs (@{code NEq} t) = term_of_fm vs (@{code NOT} (@{code Eq} t))
       
  2063   | term_of_fm vs (@{code NOT} t') = HOLogic.Not $ term_of_fm vs t'
       
  2064   | term_of_fm vs (@{code And} (t1, t2)) = HOLogic.conj $ term_of_fm vs t1 $ term_of_fm vs t2
       
  2065   | term_of_fm vs (@{code Or} (t1, t2)) = HOLogic.disj $ term_of_fm vs t1 $ term_of_fm vs t2
       
  2066   | term_of_fm vs (@{code Imp}  (t1, t2)) = HOLogic.imp $ term_of_fm vs t1 $ term_of_fm vs t2
       
  2067   | term_of_fm vs (@{code Iff} (t1, t2)) = @{term "op = :: bool \<Rightarrow> bool \<Rightarrow> bool"} $
       
  2068       term_of_fm vs t1 $ term_of_fm vs t2
       
  2069   | term_of_fm vs _ = error "If this is raised, Isabelle/HOL or generate_code is inconsistent.";
       
  2070 
       
  2071 in fn thy => fn t =>
       
  2072   let 
       
  2073     val fs = term_frees t;
       
  2074     val vs = fs ~~ (0 upto (length fs - 1));
       
  2075   in HOLogic.mk_Trueprop (HOLogic.mk_eq (t, term_of_fm vs (@{code linrqe} (fm_of_term vs t)))) end
       
  2076 end;
       
  2077 *}
       
  2078 
  2012 use "linrtac.ML"
  2079 use "linrtac.ML"
  2013 setup LinrTac.setup
  2080 setup LinrTac.setup
  2014 
  2081 
  2015 end
  2082 end