| 12869 |      1 | (*  Title:      HOL/SVC_Oracle.ML
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Lawrence C Paulson
 | 
|  |      4 |     Copyright   1999  University of Cambridge
 | 
|  |      5 | 
 | 
|  |      6 | Installing the oracle for SVC (Stanford Validity Checker)
 | 
|  |      7 | 
 | 
|  |      8 | The following code merely CALLS the oracle; 
 | 
|  |      9 |   the soundness-critical functions are at HOL/Tools/svc_funcs.ML
 | 
|  |     10 | 
 | 
|  |     11 | Based upon the work of Søren T. Heilmann
 | 
|  |     12 | *)
 | 
|  |     13 | 
 | 
|  |     14 | 
 | 
|  |     15 | (*Generalize an Isabelle formula, replacing by Vars
 | 
|  |     16 |   all subterms not intelligible to SVC.*)
 | 
|  |     17 | fun svc_abstract t =
 | 
|  |     18 |   let
 | 
|  |     19 |     (*The oracle's result is given to the subgoal using compose_tac because
 | 
|  |     20 |       its premises are matched against the assumptions rather than used
 | 
|  |     21 |       to make subgoals.  Therefore , abstraction must copy the parameters
 | 
|  |     22 |       precisely and make them available to all generated Vars.*)
 | 
|  |     23 |     val params = Term.strip_all_vars t
 | 
|  |     24 |     and body   = Term.strip_all_body t
 | 
|  |     25 |     val Us = map #2 params
 | 
|  |     26 |     val nPar = length params
 | 
|  |     27 |     val vname = ref "V_a"
 | 
|  |     28 |     val pairs = ref ([] : (term*term) list)
 | 
|  |     29 |     fun insert t = 
 | 
|  |     30 | 	let val T = fastype_of t
 | 
|  |     31 |             val v = Unify.combound (Var ((!vname,0), Us--->T),
 | 
|  |     32 | 				    0, nPar)
 | 
| 12902 |     33 | 	in  vname := Symbol.bump_string (!vname); 
 | 
| 12869 |     34 | 	    pairs := (t, v) :: !pairs;
 | 
|  |     35 | 	    v
 | 
|  |     36 | 	end;
 | 
|  |     37 |     fun replace t = 
 | 
|  |     38 | 	case t of
 | 
|  |     39 | 	    Free _  => t  (*but not existing Vars, lest the names clash*)
 | 
|  |     40 | 	  | Bound _ => t
 | 
|  |     41 | 	  | _ => (case gen_assoc Pattern.aeconv (!pairs, t) of
 | 
| 15531 |     42 | 		      SOME v => v
 | 
|  |     43 | 		    | NONE   => insert t)
 | 
| 12869 |     44 |     (*abstraction of a numeric literal*)
 | 
|  |     45 |     fun lit (t as Const("0", _)) = t
 | 
|  |     46 |       | lit (t as Const("1", _)) = t
 | 
|  |     47 |       | lit (t as Const("Numeral.number_of", _) $ w) = t
 | 
|  |     48 |       | lit t = replace t
 | 
|  |     49 |     (*abstraction of a real/rational expression*)
 | 
|  |     50 |     fun rat ((c as Const("op +", _)) $ x $ y) = c $ (rat x) $ (rat y)
 | 
|  |     51 |       | rat ((c as Const("op -", _)) $ x $ y) = c $ (rat x) $ (rat y)
 | 
|  |     52 |       | rat ((c as Const("op /", _)) $ x $ y) = c $ (rat x) $ (rat y)
 | 
|  |     53 |       | rat ((c as Const("op *", _)) $ x $ y) = c $ (rat x) $ (rat y)
 | 
|  |     54 |       | rat ((c as Const("uminus", _)) $ x) = c $ (rat x)
 | 
|  |     55 |       | rat t = lit t
 | 
|  |     56 |     (*abstraction of an integer expression: no div, mod*)
 | 
|  |     57 |     fun int ((c as Const("op +", _)) $ x $ y) = c $ (int x) $ (int y)
 | 
|  |     58 |       | int ((c as Const("op -", _)) $ x $ y) = c $ (int x) $ (int y)
 | 
|  |     59 |       | int ((c as Const("op *", _)) $ x $ y) = c $ (int x) $ (int y)
 | 
|  |     60 |       | int ((c as Const("uminus", _)) $ x) = c $ (int x)
 | 
|  |     61 |       | int t = lit t
 | 
|  |     62 |     (*abstraction of a natural number expression: no minus*)
 | 
|  |     63 |     fun nat ((c as Const("op +", _)) $ x $ y) = c $ (nat x) $ (nat y)
 | 
|  |     64 |       | nat ((c as Const("op *", _)) $ x $ y) = c $ (nat x) $ (nat y)
 | 
|  |     65 |       | nat ((c as Const("Suc", _)) $ x) = c $ (nat x)
 | 
|  |     66 |       | nat t = lit t
 | 
|  |     67 |     (*abstraction of a relation: =, <, <=*)
 | 
|  |     68 |     fun rel (T, c $ x $ y) =
 | 
|  |     69 | 	    if T = HOLogic.realT then c $ (rat x) $ (rat y)
 | 
|  |     70 | 	    else if T = HOLogic.intT then c $ (int x) $ (int y)
 | 
|  |     71 | 	    else if T = HOLogic.natT then c $ (nat x) $ (nat y)
 | 
|  |     72 | 	    else if T = HOLogic.boolT then c $ (fm x) $ (fm y)
 | 
|  |     73 | 	    else replace (c $ x $ y)   (*non-numeric comparison*)
 | 
|  |     74 |     (*abstraction of a formula*)
 | 
|  |     75 |     and fm ((c as Const("op &", _)) $ p $ q) = c $ (fm p) $ (fm q)
 | 
|  |     76 |       | fm ((c as Const("op |", _)) $ p $ q) = c $ (fm p) $ (fm q)
 | 
|  |     77 |       | fm ((c as Const("op -->", _)) $ p $ q) = c $ (fm p) $ (fm q)
 | 
|  |     78 |       | fm ((c as Const("Not", _)) $ p) = c $ (fm p)
 | 
|  |     79 |       | fm ((c as Const("True", _))) = c
 | 
|  |     80 |       | fm ((c as Const("False", _))) = c
 | 
|  |     81 |       | fm (t as Const("op =",  Type ("fun", [T,_])) $ _ $ _) = rel (T, t)
 | 
|  |     82 |       | fm (t as Const("op <",  Type ("fun", [T,_])) $ _ $ _) = rel (T, t)
 | 
|  |     83 |       | fm (t as Const("op <=", Type ("fun", [T,_])) $ _ $ _) = rel (T, t)
 | 
|  |     84 |       | fm t = replace t
 | 
|  |     85 |     (*entry point, and abstraction of a meta-formula*)
 | 
|  |     86 |     fun mt ((c as Const("Trueprop", _)) $ p) = c $ (fm p)
 | 
|  |     87 |       | mt ((c as Const("==>", _)) $ p $ q)  = c $ (mt p) $ (mt q)
 | 
|  |     88 |       | mt t = fm t  (*it might be a formula*)
 | 
|  |     89 |   in (list_all (params, mt body), !pairs) end;
 | 
|  |     90 | 
 | 
|  |     91 | 
 | 
|  |     92 | (*Present the entire subgoal to the oracle, assumptions and all, but possibly
 | 
|  |     93 |   abstracted.  Use via compose_tac, which performs no lifting but will
 | 
|  |     94 |   instantiate variables.*)
 | 
|  |     95 | local val svc_thy = the_context () in
 | 
|  |     96 | 
 | 
|  |     97 | fun svc_tac i st = 
 | 
|  |     98 |   let val prem = BasisLibrary.List.nth (prems_of st, i-1)
 | 
|  |     99 |       val (absPrem, _) = svc_abstract prem
 | 
|  |    100 |       val th = invoke_oracle svc_thy "svc_oracle"
 | 
|  |    101 | 	             (#sign (rep_thm st), Svc.OracleExn absPrem)
 | 
|  |    102 |    in 
 | 
|  |    103 |       compose_tac (false, th, 0) i st
 | 
|  |    104 |    end 
 | 
|  |    105 |    handle Svc.OracleExn _ => Seq.empty
 | 
|  |    106 | 	| Subscript       => Seq.empty;
 | 
|  |    107 | 
 | 
|  |    108 | end;
 | 
|  |    109 | 
 | 
|  |    110 | 
 | 
|  |    111 | (*check if user has SVC installed*)
 | 
|  |    112 | fun svc_enabled () = getenv "SVC_HOME" <> "";
 | 
|  |    113 | fun if_svc_enabled f x = if svc_enabled () then f x else ();
 |