| author | wenzelm | 
| Tue, 05 Oct 1999 18:26:34 +0200 | |
| changeset 7742 | 01386eb4eab0 | 
| parent 7545 | 1578f1fd62cf | 
| child 10892 | 405b077433a3 | 
| permissions | -rw-r--r-- | 
| 7145 | 1 | (* Title: HOL/Tools/svc_funcs.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson | |
| 4 | Copyright 1999 University of Cambridge | |
| 5 | ||
| 7285 | 6 | Translation functions for the interface to SVC | 
| 7145 | 7 | |
| 8 | Based upon the work of Søren T. Heilmann | |
| 9 | ||
| 7164 | 10 | Integers and naturals are translated as follows: | 
| 11 | In a positive context, replace x<y by x+1<=y | |
| 12 | In a negative context, replace x<=y by x<y+1 | |
| 13 | In a negative context, replace x=y by x<y+1 & y<x+1 | |
| 14 | Biconditionals (if-and-only-iff) are expanded if they require such translations | |
| 15 | in either operand. | |
| 7145 | 16 | |
| 7164 | 17 | For each variable of type nat, an assumption is added that it is non-negative. | 
| 18 | *) | |
| 7145 | 19 | |
| 20 | structure Svc = | |
| 21 | struct | |
| 22 | val trace = ref false; | |
| 23 | ||
| 24 | datatype expr = | |
| 7285 | 25 | Buildin of string * expr list | 
| 26 | | Interp of string * expr list | |
| 27 | | UnInterp of string * expr list | |
| 28 | | FalseExpr | |
| 29 | | TrueExpr | |
| 30 | | Int of int | |
| 31 | | Rat of int * int; | |
| 7145 | 32 | |
| 33 | open BasisLibrary | |
| 34 | ||
| 7285 | 35 | fun signedInt i = | 
| 36 | if i < 0 then "-" ^ Int.toString (~i) | |
| 37 | else Int.toString i; | |
| 38 | ||
| 39 | fun is_intnat T = T = HOLogic.intT orelse T = HOLogic.natT; | |
| 40 | ||
| 41 | fun is_numeric T = is_intnat T orelse T = HOLogic.realT; | |
| 42 | ||
| 43 | fun is_numeric_op T = is_numeric (domain_type T); | |
| 44 | ||
| 7145 | 45 | fun toString t = | 
| 7285 | 46 | let fun ue (Buildin(s, l)) = | 
| 7145 | 47 | 	     "(" ^ s ^ (foldl (fn (a, b) => a ^ " " ^ (ue b)) ("", l)) ^ ") "
 | 
| 7285 | 48 | | ue (Interp(s, l)) = | 
| 7145 | 49 | 	     "{" ^ s ^ (foldl (fn (a, b) => a ^ " " ^ (ue b)) ("", l)) ^ "} "
 | 
| 7285 | 50 | | ue (UnInterp(s, l)) = | 
| 7145 | 51 | 	     "(" ^ s ^ (foldl (fn (a, b) => a ^ " " ^ (ue b)) ("", l)) ^ ") "
 | 
| 7285 | 52 | | ue (FalseExpr) = "FALSE " | 
| 53 | | ue (TrueExpr) = "TRUE " | |
| 54 | | ue (Int i) = (signedInt i) ^ " " | |
| 55 | | ue (Rat(i, j)) = (signedInt i) ^ "|" ^ (signedInt j) ^ " " | |
| 7145 | 56 | in | 
| 57 | ue t | |
| 58 | end; | |
| 59 | ||
| 60 | fun valid e = | |
| 61 | let val svc_home = getenv "SVC_HOME" | |
| 62 | val svc_machine = getenv "SVC_MACHINE" | |
| 63 | val check_valid = if svc_home = "" | |
| 64 | then error "Environment variable SVC_HOME not set" | |
| 65 | else if svc_machine = "" | |
| 66 | then error "Environment variable SVC_MACHINE not set" | |
| 67 | else svc_home ^ "/" ^ svc_machine ^ "/bin/check_valid" | |
| 68 | val svc_input = toString e | |
| 69 |       val _ = if !trace then writeln ("Calling SVC:\n" ^ svc_input) else ()
 | |
| 70 | val svc_input_file = File.tmp_path (Path.basic "SVM_in"); | |
| 71 | val svc_output_file = File.tmp_path (Path.basic "SVM_out"); | |
| 72 | val _ = (File.write svc_input_file svc_input; | |
| 73 | execute (check_valid ^ " -dump-result " ^ | |
| 7164 | 74 | File.sysify_path svc_output_file ^ | 
| 75 | " " ^ File.sysify_path svc_input_file ^ | |
| 7145 | 76 | "> /dev/null 2>&1")) | 
| 77 | val svc_output = File.read svc_output_file | |
| 78 | handle _ => error "SVC returned no output" | |
| 79 | in | |
| 7285 | 80 |       if ! trace then writeln ("SVC Returns:\n" ^ svc_output)
 | 
| 81 | else (File.rm svc_input_file; File.rm svc_output_file); | |
| 7145 | 82 | String.isPrefix "VALID" svc_output | 
| 83 | end | |
| 84 | ||
| 85 | (*New exception constructor for passing arguments to the oracle*) | |
| 86 | exception OracleExn of term; | |
| 87 | ||
| 7164 | 88 | fun apply c args = | 
| 89 | let val (ts, bs) = ListPair.unzip args | |
| 90 | in (list_comb(c,ts), exists I bs) end; | |
| 91 | ||
| 7285 | 92 | (*Determining whether the biconditionals must be unfolded: if there are | 
| 7164 | 93 | int or nat comparisons below*) | 
| 94 | val iff_tag = | |
| 95 | let fun tag t = | |
| 96 | let val (c,ts) = strip_comb t | |
| 97 | in case c of | |
| 98 | 	     Const("op &", _)   => apply c (map tag ts)
 | |
| 99 | 	   | Const("op |", _)   => apply c (map tag ts)
 | |
| 100 | 	   | Const("op -->", _) => apply c (map tag ts)
 | |
| 101 | 	   | Const("Not", _)    => apply c (map tag ts)
 | |
| 102 | 	   | Const("True", _)   => (c, false)
 | |
| 103 | 	   | Const("False", _)  => (c, false)
 | |
| 104 | 	   | Const("op =", Type ("fun", [T,_])) => 
 | |
| 105 | if T = HOLogic.boolT then | |
| 106 | (*biconditional: with int/nat comparisons below?*) | |
| 107 | let val [t1,t2] = ts | |
| 108 | val (u1,b1) = tag t1 | |
| 109 | and (u2,b2) = tag t2 | |
| 110 | val cname = if b1 orelse b2 then "unfold" else "keep" | |
| 111 | in | |
| 112 | 			(Const ("SVC_Oracle.iff_" ^ cname, dummyT) $ u1 $ u2,
 | |
| 113 | b1 orelse b2) | |
| 114 | end | |
| 7285 | 115 | else (*might be numeric equality*) (t, is_intnat T) | 
| 7164 | 116 | 	   | Const("op <", Type ("fun", [T,_]))  => (t, is_intnat T)
 | 
| 117 | 	   | Const("op <=", Type ("fun", [T,_])) => (t, is_intnat T)
 | |
| 118 | | _ => (t, false) | |
| 119 | end | |
| 120 | in #1 o tag end; | |
| 121 | ||
| 122 | (*Map expression e to 0<=a --> e, where "a" is the name of a nat variable*) | |
| 123 | fun add_nat_var (a, e) = | |
| 7285 | 124 |      Buildin("=>", [Buildin("<=", [Int 0, UnInterp (a, [])]),
 | 
| 125 | e]); | |
| 7164 | 126 | |
| 7545 
1578f1fd62cf
fixed SOUNDNESS BUG concerning the map from terms like ?f x y to SVC variables
 paulson parents: 
7285diff
changeset | 127 | fun param_string [] = "" | 
| 
1578f1fd62cf
fixed SOUNDNESS BUG concerning the map from terms like ?f x y to SVC variables
 paulson parents: 
7285diff
changeset | 128 | | param_string is = "_" ^ space_implode "_" (map string_of_int is) | 
| 
1578f1fd62cf
fixed SOUNDNESS BUG concerning the map from terms like ?f x y to SVC variables
 paulson parents: 
7285diff
changeset | 129 | |
| 7145 | 130 | (*Translate an Isabelle formula into an SVC expression | 
| 131 | pos ["positive"]: true if an assumption, false if a goal*) | |
| 132 | fun expr_of pos t = | |
| 133 | let | |
| 7164 | 134 | val params = rev (rename_wrt_term t (Term.strip_all_vars t)) | 
| 135 | and body = Term.strip_all_body t | |
| 136 | val nat_vars = ref ([] : string list) | |
| 137 | (*translation of a variable: record all natural numbers*) | |
| 7545 
1578f1fd62cf
fixed SOUNDNESS BUG concerning the map from terms like ?f x y to SVC variables
 paulson parents: 
7285diff
changeset | 138 | fun trans_var (a,T,is) = | 
| 7164 | 139 | (if T = HOLogic.natT then nat_vars := (a ins_string (!nat_vars)) | 
| 140 | else (); | |
| 7545 
1578f1fd62cf
fixed SOUNDNESS BUG concerning the map from terms like ?f x y to SVC variables
 paulson parents: 
7285diff
changeset | 141 | UnInterp (a ^ param_string is, [])) | 
| 
1578f1fd62cf
fixed SOUNDNESS BUG concerning the map from terms like ?f x y to SVC variables
 paulson parents: 
7285diff
changeset | 142 | (*A variable, perhaps applied to a series of parameters*) | 
| 
1578f1fd62cf
fixed SOUNDNESS BUG concerning the map from terms like ?f x y to SVC variables
 paulson parents: 
7285diff
changeset | 143 |     fun var (Free(a,T), is)      = trans_var ("F_" ^ a, T, is)
 | 
| 
1578f1fd62cf
fixed SOUNDNESS BUG concerning the map from terms like ?f x y to SVC variables
 paulson parents: 
7285diff
changeset | 144 | | var (Var((a, 0), T), is) = trans_var (a, T, is) | 
| 
1578f1fd62cf
fixed SOUNDNESS BUG concerning the map from terms like ?f x y to SVC variables
 paulson parents: 
7285diff
changeset | 145 | | var (Bound i, is) = | 
| 7164 | 146 | let val (a,T) = List.nth (params, i) | 
| 7545 
1578f1fd62cf
fixed SOUNDNESS BUG concerning the map from terms like ?f x y to SVC variables
 paulson parents: 
7285diff
changeset | 147 | 	  in  trans_var ("B_" ^ a, T, is)  end
 | 
| 
1578f1fd62cf
fixed SOUNDNESS BUG concerning the map from terms like ?f x y to SVC variables
 paulson parents: 
7285diff
changeset | 148 | | var (t $ Bound i, is) = var(t,i::is) | 
| 
1578f1fd62cf
fixed SOUNDNESS BUG concerning the map from terms like ?f x y to SVC variables
 paulson parents: 
7285diff
changeset | 149 | (*removing a parameter from a Var: the bound var index will | 
| 
1578f1fd62cf
fixed SOUNDNESS BUG concerning the map from terms like ?f x y to SVC variables
 paulson parents: 
7285diff
changeset | 150 | become part of the Var's name*) | 
| 
1578f1fd62cf
fixed SOUNDNESS BUG concerning the map from terms like ?f x y to SVC variables
 paulson parents: 
7285diff
changeset | 151 | | var (t,_) = raise OracleExn t; | 
| 7164 | 152 | (*translation of a literal*) | 
| 153 |     fun lit (Const("Numeral.number_of", _) $ w) = NumeralSyntax.dest_bin w
 | |
| 154 |       | lit (Const("0", _)) = 0
 | |
| 155 |       | lit (Const("RealDef.0r", _)) = 0
 | |
| 156 |       | lit (Const("RealDef.1r", _)) = 1
 | |
| 157 | (*translation of a literal expression [no variables]*) | |
| 7285 | 158 |     fun litExp (Const("op +", T) $ x $ y) = 
 | 
| 159 | if is_numeric_op T then (litExp x) + (litExp y) | |
| 160 | else raise OracleExn t | |
| 161 |       | litExp (Const("op -", T) $ x $ y) = 
 | |
| 162 | if is_numeric_op T then (litExp x) - (litExp y) | |
| 163 | else raise OracleExn t | |
| 164 |       | litExp (Const("op *", T) $ x $ y) = 
 | |
| 165 | if is_numeric_op T then (litExp x) * (litExp y) | |
| 166 | else raise OracleExn t | |
| 167 |       | litExp (Const("uminus", T) $ x)   = 
 | |
| 168 | if is_numeric_op T then ~(litExp x) | |
| 169 | else raise OracleExn t | |
| 7164 | 170 | | litExp t = lit t | 
| 7285 | 171 | handle Match => raise OracleExn t | 
| 7164 | 172 | (*translation of a real/rational expression*) | 
| 7285 | 173 |     fun suc t = Interp("+", [Int 1, t])
 | 
| 7164 | 174 |     fun tm (Const("Suc", T) $ x) = suc (tm x)
 | 
| 7285 | 175 |       | tm (Const("op +", T) $ x $ y) = 
 | 
| 176 | 	  if is_numeric_op T then Interp("+", [tm x, tm y])
 | |
| 177 | else raise OracleExn t | |
| 178 |       | tm (Const("op -", T) $ x $ y) = 
 | |
| 179 | if is_numeric_op T then | |
| 180 | 	      Interp("+", [tm x, Interp("*", [Int ~1, tm y])])
 | |
| 181 | else raise OracleExn t | |
| 182 |       | tm (Const("op *", T) $ x $ y) = 
 | |
| 183 | 	  if is_numeric_op T then Interp("*", [tm x, tm y])
 | |
| 184 | else raise OracleExn t | |
| 185 |       | tm (Const("RealDef.rinv", T) $ x) = 
 | |
| 186 | if domain_type T = HOLogic.realT then | |
| 187 | Rat(1, litExp x) | |
| 188 | else raise OracleExn t | |
| 189 |       | tm (Const("uminus", T) $ x) = 
 | |
| 190 | 	  if is_numeric_op T then Interp("*", [Int ~1, tm x])
 | |
| 191 | else raise OracleExn t | |
| 192 | | tm t = Int (lit t) | |
| 7545 
1578f1fd62cf
fixed SOUNDNESS BUG concerning the map from terms like ?f x y to SVC variables
 paulson parents: 
7285diff
changeset | 193 | handle Match => var (t,[]) | 
| 7164 | 194 | (*translation of a formula*) | 
| 195 |     and fm pos (Const("op &", _) $ p $ q) =  
 | |
| 7285 | 196 | 	    Buildin("AND", [fm pos p, fm pos q])
 | 
| 7164 | 197 |       | fm pos (Const("op |", _) $ p $ q) =  
 | 
| 7285 | 198 | 	    Buildin("OR", [fm pos p, fm pos q])
 | 
| 7164 | 199 |       | fm pos (Const("op -->", _) $ p $ q) =  
 | 
| 7285 | 200 | 	    Buildin("=>", [fm (not pos) p, fm pos q])
 | 
| 7164 | 201 |       | fm pos (Const("Not", _) $ p) =  
 | 
| 7285 | 202 | 	    Buildin("NOT", [fm (not pos) p])
 | 
| 203 |       | fm pos (Const("True", _)) = TrueExpr
 | |
| 204 |       | fm pos (Const("False", _)) = FalseExpr
 | |
| 7164 | 205 |       | fm pos (Const("SVC_Oracle.iff_keep", _) $ p $ q) = 
 | 
| 206 | (*polarity doesn't matter*) | |
| 7285 | 207 | 	    Buildin("=", [fm pos p, fm pos q]) 
 | 
| 7164 | 208 |       | fm pos (Const("SVC_Oracle.iff_unfold", _) $ p $ q) = 
 | 
| 7285 | 209 | 	    Buildin("AND",   (*unfolding uses both polarities*)
 | 
| 210 | 			 [Buildin("=>", [fm (not pos) p, fm pos q]),
 | |
| 211 | 			  Buildin("=>", [fm (not pos) q, fm pos p])])
 | |
| 7164 | 212 |       | fm pos (t as Const("op =", Type ("fun", [T,_])) $ x $ y) = 
 | 
| 213 | let val tx = tm x and ty = tm y | |
| 214 | in if pos orelse T = HOLogic.realT then | |
| 7285 | 215 | 		       Buildin("=", [tx, ty])
 | 
| 7164 | 216 | else if is_intnat T then | 
| 7285 | 217 | 		       Buildin("AND", 
 | 
| 218 | 				    [Buildin("<", [tx, suc ty]), 
 | |
| 219 | 				     Buildin("<", [ty, suc tx])])
 | |
| 7164 | 220 | else raise OracleExn t | 
| 221 | end | |
| 222 | (*inequalities: possible types are nat, int, real*) | |
| 223 |       | fm pos (t as Const("op <",  Type ("fun", [T,_])) $ x $ y) = 
 | |
| 224 | if not pos orelse T = HOLogic.realT then | |
| 7285 | 225 | 		Buildin("<", [tm x, tm y])
 | 
| 7164 | 226 | else if is_intnat T then | 
| 7285 | 227 | 		Buildin("<=", [suc (tm x), tm y])
 | 
| 7164 | 228 | else raise OracleExn t | 
| 229 |       | fm pos (t as Const("op <=",  Type ("fun", [T,_])) $ x $ y) = 
 | |
| 230 | if pos orelse T = HOLogic.realT then | |
| 7285 | 231 | 		Buildin("<=", [tm x, tm y])
 | 
| 7164 | 232 | else if is_intnat T then | 
| 7285 | 233 | 		Buildin("<", [tm x, suc (tm y)])
 | 
| 7164 | 234 | else raise OracleExn t | 
| 7545 
1578f1fd62cf
fixed SOUNDNESS BUG concerning the map from terms like ?f x y to SVC variables
 paulson parents: 
7285diff
changeset | 235 | | fm pos t = var(t,[]); | 
| 7164 | 236 | (*entry point, and translation of a meta-formula*) | 
| 237 |       fun mt pos ((c as Const("Trueprop", _)) $ p) = fm pos (iff_tag p)
 | |
| 238 | 	| mt pos ((c as Const("==>", _)) $ p $ q) = 
 | |
| 7285 | 239 | 	    Buildin("=>", [mt (not pos) p, mt pos q])
 | 
| 7164 | 240 | | mt pos t = fm pos (iff_tag t) (*it might be a formula*) | 
| 241 | ||
| 242 | val body_e = mt pos body (*evaluate now to assign into !nat_vars*) | |
| 7145 | 243 | in | 
| 7164 | 244 | foldr add_nat_var (!nat_vars, body_e) | 
| 7145 | 245 | end; | 
| 246 | ||
| 247 | ||
| 7285 | 248 | (*The oracle proves the given formula t, if possible*) | 
| 249 | fun oracle (sign, OracleExn t) = | |
| 250 |    let val dummy = if !trace then writeln ("Subgoal abstracted to\n" ^
 | |
| 251 | Sign.string_of_term sign t) | |
| 7164 | 252 | else () | 
| 253 | in | |
| 7285 | 254 | if valid (expr_of false t) then t | 
| 255 | else raise OracleExn t | |
| 7164 | 256 | end; | 
| 7145 | 257 | |
| 258 | end; |