| author | wenzelm | 
| Mon, 15 Jul 2013 19:23:19 +0200 | |
| changeset 52663 | 6e71d43775e5 | 
| parent 51940 | 958d439b3013 | 
| child 56245 | 84fc7dfa3cd4 | 
| permissions | -rw-r--r-- | 
| 24584 | 1 | (* Title: HOL/ex/svc_funcs.ML | 
| 12869 | 2 | Author: Lawrence C Paulson | 
| 3 | Copyright 1999 University of Cambridge | |
| 4 | ||
| 29276 | 5 | Translation functions for the interface to SVC. | 
| 12869 | 6 | |
| 19336 
fb5e19d26d5e
removed some illegal characters: they were crashing SML/NJ
 paulson parents: 
19277diff
changeset | 7 | Based upon the work of Soren T. Heilmann | 
| 12869 | 8 | |
| 9 | Integers and naturals are translated as follows: | |
| 10 | In a positive context, replace x<y by x+1<=y | |
| 11 | In a negative context, replace x<=y by x<y+1 | |
| 12 | In a negative context, replace x=y by x<y+1 & y<x+1 | |
| 13 | Biconditionals (if-and-only-iff) are expanded if they require such translations | |
| 14 | in either operand. | |
| 15 | ||
| 16 | For each variable of type nat, an assumption is added that it is non-negative. | |
| 51940 | 17 | |
| 18 | Relevant Isabelle environment settings: | |
| 19 | ||
| 20 | #SVC_HOME= | |
| 21 | #SVC_MACHINE=i386-redhat-linux | |
| 22 | #SVC_MACHINE=sparc-sun-solaris | |
| 12869 | 23 | *) | 
| 24 | ||
| 25 | structure Svc = | |
| 26 | struct | |
| 32740 | 27 | val trace = Unsynchronized.ref false; | 
| 12869 | 28 | |
| 29 | datatype expr = | |
| 30 | Buildin of string * expr list | |
| 31 | | Interp of string * expr list | |
| 32 | | UnInterp of string * expr list | |
| 16836 | 33 | | FalseExpr | 
| 12869 | 34 | | TrueExpr | 
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24584diff
changeset | 35 | | Int of int | 
| 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24584diff
changeset | 36 | | Rat of int * int; | 
| 16836 | 37 | |
| 12869 | 38 | fun is_intnat T = T = HOLogic.intT orelse T = HOLogic.natT; | 
| 16836 | 39 | |
| 12869 | 40 | fun is_numeric T = is_intnat T orelse T = HOLogic.realT; | 
| 16836 | 41 | |
| 12869 | 42 | fun is_numeric_op T = is_numeric (domain_type T); | 
| 43 | ||
| 44 | fun toString t = | |
| 16836 | 45 | let fun ue (Buildin(s, l)) = | 
| 46 |              "(" ^ s ^ (Library.foldl (fn (a, b) => a ^ " " ^ (ue b)) ("", l)) ^ ") "
 | |
| 47 | | ue (Interp(s, l)) = | |
| 48 |              "{" ^ s ^ (Library.foldl (fn (a, b) => a ^ " " ^ (ue b)) ("", l)) ^ "} "
 | |
| 49 | | ue (UnInterp(s, l)) = | |
| 50 |              "(" ^ s ^ (Library.foldl (fn (a, b) => a ^ " " ^ (ue b)) ("", l)) ^ ") "
 | |
| 51 | | ue (FalseExpr) = "FALSE " | |
| 52 | | ue (TrueExpr) = "TRUE " | |
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24584diff
changeset | 53 | | ue (Int i) = signed_string_of_int i ^ " " | 
| 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24584diff
changeset | 54 | | ue (Rat(i, j)) = signed_string_of_int i ^ "|" ^ signed_string_of_int j ^ " " | 
| 12869 | 55 | in | 
| 16836 | 56 | ue t | 
| 12869 | 57 | end; | 
| 58 | ||
| 16836 | 59 | fun valid e = | 
| 60 | let val svc_home = getenv "SVC_HOME" | |
| 12869 | 61 | val svc_machine = getenv "SVC_MACHINE" | 
| 62 | val check_valid = if svc_home = "" | |
| 16836 | 63 | then error "Environment variable SVC_HOME not set" | 
| 64 | else if svc_machine = "" | |
| 65 | then error "Environment variable SVC_MACHINE not set" | |
| 66 | else svc_home ^ "/" ^ svc_machine ^ "/bin/check_valid" | |
| 12869 | 67 | val svc_input = toString e | 
| 68 |       val _ = if !trace then tracing ("Calling SVC:\n" ^ svc_input) else ()
 | |
| 69 | val svc_input_file = File.tmp_path (Path.basic "SVM_in"); | |
| 70 | val svc_output_file = File.tmp_path (Path.basic "SVM_out"); | |
| 35010 
d6e492cea6e4
renamed system/system_out to bash/bash_output -- to emphasized that this is really GNU bash, not some undefined POSIX sh;
 wenzelm parents: 
34974diff
changeset | 71 | val _ = File.write svc_input_file svc_input; | 
| 
d6e492cea6e4
renamed system/system_out to bash/bash_output -- to emphasized that this is really GNU bash, not some undefined POSIX sh;
 wenzelm parents: 
34974diff
changeset | 72 | val _ = | 
| 43850 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 wenzelm parents: 
42364diff
changeset | 73 | Isabelle_System.bash_output (check_valid ^ " -dump-result " ^ | 
| 35010 
d6e492cea6e4
renamed system/system_out to bash/bash_output -- to emphasized that this is really GNU bash, not some undefined POSIX sh;
 wenzelm parents: 
34974diff
changeset | 74 | File.shell_path svc_output_file ^ " " ^ File.shell_path svc_input_file ^ | 
| 
d6e492cea6e4
renamed system/system_out to bash/bash_output -- to emphasized that this is really GNU bash, not some undefined POSIX sh;
 wenzelm parents: 
34974diff
changeset | 75 | ">/dev/null 2>&1") | 
| 12869 | 76 | val svc_output = | 
| 21395 
f34ac19659ae
moved some fundamental concepts to General/basics.ML;
 wenzelm parents: 
21078diff
changeset | 77 | (case try File.read svc_output_file of | 
| 15531 | 78 | SOME out => out | 
| 79 | | NONE => error "SVC returned no output"); | |
| 12869 | 80 | in | 
| 81 |       if ! trace then tracing ("SVC Returns:\n" ^ svc_output)
 | |
| 82 | else (File.rm svc_input_file; File.rm svc_output_file); | |
| 83 | String.isPrefix "VALID" svc_output | |
| 84 | end | |
| 85 | ||
| 16836 | 86 |  fun fail t = raise TERM ("SVC oracle", [t]);
 | 
| 12869 | 87 | |
| 88 | fun apply c args = | |
| 89 | let val (ts, bs) = ListPair.unzip args | |
| 90 | in (list_comb(c,ts), exists I bs) end; | |
| 91 | ||
| 92 | (*Determining whether the biconditionals must be unfolded: if there are | |
| 93 | int or nat comparisons below*) | |
| 94 | val iff_tag = | |
| 95 | let fun tag t = | |
| 16836 | 96 | let val (c,ts) = strip_comb t | 
| 97 | in case c of | |
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 98 |              Const(@{const_name HOL.conj}, _)   => apply c (map tag ts)
 | 
| 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 99 |            | Const(@{const_name HOL.disj}, _)   => apply c (map tag ts)
 | 
| 38786 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38558diff
changeset | 100 |            | Const(@{const_name HOL.implies}, _) => apply c (map tag ts)
 | 
| 38558 | 101 |            | Const(@{const_name Not}, _)    => apply c (map tag ts)
 | 
| 102 |            | Const(@{const_name True}, _)   => (c, false)
 | |
| 103 |            | Const(@{const_name False}, _)  => (c, false)
 | |
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 104 |            | Const(@{const_name HOL.eq}, Type ("fun", [T,_])) =>
 | 
| 16836 | 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 | |
| 115 | else (*might be numeric equality*) (t, is_intnat T) | |
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 116 |            | Const(@{const_name Orderings.less}, Type ("fun", [T,_]))  => (t, is_intnat T)
 | 
| 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 117 |            | Const(@{const_name Orderings.less_eq}, Type ("fun", [T,_])) => (t, is_intnat T)
 | 
| 16836 | 118 | | _ => (t, false) | 
| 119 | end | |
| 12869 | 120 | in #1 o tag end; | 
| 121 | ||
| 122 | (*Map expression e to 0<=a --> e, where "a" is the name of a nat variable*) | |
| 21078 | 123 | fun add_nat_var a e = | 
| 12869 | 124 |      Buildin("=>", [Buildin("<=", [Int 0, UnInterp (a, [])]),
 | 
| 16836 | 125 | e]); | 
| 12869 | 126 | |
| 127 | fun param_string [] = "" | |
| 128 | | param_string is = "_" ^ space_implode "_" (map string_of_int is) | |
| 129 | ||
| 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 | |
| 29276 | 134 | val params = rev (Term.rename_wrt_term t (Term.strip_all_vars t)) | 
| 12869 | 135 | and body = Term.strip_all_body t | 
| 32740 | 136 | val nat_vars = Unsynchronized.ref ([] : string list) | 
| 12869 | 137 | (*translation of a variable: record all natural numbers*) | 
| 138 | fun trans_var (a,T,is) = | |
| 20853 | 139 | (if T = HOLogic.natT then nat_vars := (insert (op =) a (!nat_vars)) | 
| 16836 | 140 | else (); | 
| 12869 | 141 | UnInterp (a ^ param_string is, [])) | 
| 142 | (*A variable, perhaps applied to a series of parameters*) | |
| 143 |     fun var (Free(a,T), is)      = trans_var ("F_" ^ a, T, is)
 | |
| 144 | | var (Var((a, 0), T), is) = trans_var (a, T, is) | |
| 16836 | 145 | | var (Bound i, is) = | 
| 42364 | 146 | let val (a,T) = nth params i | 
| 16836 | 147 |           in  trans_var ("B_" ^ a, T, is)  end
 | 
| 12869 | 148 | | var (t $ Bound i, is) = var(t,i::is) | 
| 149 | (*removing a parameter from a Var: the bound var index will | |
| 150 | become part of the Var's name*) | |
| 16836 | 151 | | var (t,_) = fail t; | 
| 12869 | 152 | (*translation of a literal*) | 
| 21820 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 haftmann parents: 
21395diff
changeset | 153 | val lit = snd o HOLogic.dest_number; | 
| 12869 | 154 | (*translation of a literal expression [no variables]*) | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 155 |     fun litExp (Const(@{const_name Groups.plus}, T) $ x $ y) =
 | 
| 16836 | 156 | if is_numeric_op T then (litExp x) + (litExp y) | 
| 157 | else fail t | |
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 158 |       | litExp (Const(@{const_name Groups.minus}, T) $ x $ y) =
 | 
| 16836 | 159 | if is_numeric_op T then (litExp x) - (litExp y) | 
| 160 | else fail t | |
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 161 |       | litExp (Const(@{const_name Groups.times}, T) $ x $ y) =
 | 
| 16836 | 162 | if is_numeric_op T then (litExp x) * (litExp y) | 
| 163 | else fail t | |
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 164 |       | litExp (Const(@{const_name Groups.uminus}, T) $ x)   =
 | 
| 16836 | 165 | if is_numeric_op T then ~(litExp x) | 
| 166 | else fail t | |
| 167 | | litExp t = lit t | |
| 168 | handle Match => fail t | |
| 12869 | 169 | (*translation of a real/rational expression*) | 
| 170 |     fun suc t = Interp("+", [Int 1, t])
 | |
| 22997 | 171 |     fun tm (Const(@{const_name Suc}, T) $ x) = suc (tm x)
 | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 172 |       | tm (Const(@{const_name Groups.plus}, T) $ x $ y) =
 | 
| 16836 | 173 |           if is_numeric_op T then Interp("+", [tm x, tm y])
 | 
| 174 | else fail t | |
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 175 |       | tm (Const(@{const_name Groups.minus}, T) $ x $ y) =
 | 
| 16836 | 176 | if is_numeric_op T then | 
| 177 |               Interp("+", [tm x, Interp("*", [Int ~1, tm y])])
 | |
| 178 | else fail t | |
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 179 |       | tm (Const(@{const_name Groups.times}, T) $ x $ y) =
 | 
| 16836 | 180 |           if is_numeric_op T then Interp("*", [tm x, tm y])
 | 
| 181 | else fail t | |
| 44064 
5bce8ff0d9ae
moved division ring stuff from Rings.thy to Fields.thy
 huffman parents: 
43850diff
changeset | 182 |       | tm (Const(@{const_name Fields.inverse}, T) $ x) =
 | 
| 16836 | 183 | if domain_type T = HOLogic.realT then | 
| 184 | Rat(1, litExp x) | |
| 185 | else fail t | |
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 186 |       | tm (Const(@{const_name Groups.uminus}, T) $ x) =
 | 
| 16836 | 187 |           if is_numeric_op T then Interp("*", [Int ~1, tm x])
 | 
| 188 | else fail t | |
| 189 | | tm t = Int (lit t) | |
| 190 | handle Match => var (t,[]) | |
| 12869 | 191 | (*translation of a formula*) | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 192 |     and fm pos (Const(@{const_name HOL.conj}, _) $ p $ q) =
 | 
| 16836 | 193 |             Buildin("AND", [fm pos p, fm pos q])
 | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 194 |       | fm pos (Const(@{const_name HOL.disj}, _) $ p $ q) =
 | 
| 16836 | 195 |             Buildin("OR", [fm pos p, fm pos q])
 | 
| 38786 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38558diff
changeset | 196 |       | fm pos (Const(@{const_name HOL.implies}, _) $ p $ q) =
 | 
| 16836 | 197 |             Buildin("=>", [fm (not pos) p, fm pos q])
 | 
| 38558 | 198 |       | fm pos (Const(@{const_name Not}, _) $ p) =
 | 
| 16836 | 199 |             Buildin("NOT", [fm (not pos) p])
 | 
| 38558 | 200 |       | fm pos (Const(@{const_name True}, _)) = TrueExpr
 | 
| 201 |       | fm pos (Const(@{const_name False}, _)) = FalseExpr
 | |
| 16836 | 202 |       | fm pos (Const("SVC_Oracle.iff_keep", _) $ p $ q) =
 | 
| 203 | (*polarity doesn't matter*) | |
| 204 |             Buildin("=", [fm pos p, fm pos q])
 | |
| 205 |       | fm pos (Const("SVC_Oracle.iff_unfold", _) $ p $ q) =
 | |
| 206 |             Buildin("AND",   (*unfolding uses both polarities*)
 | |
| 207 |                          [Buildin("=>", [fm (not pos) p, fm pos q]),
 | |
| 208 |                           Buildin("=>", [fm (not pos) q, fm pos p])])
 | |
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 209 |       | fm pos (t as Const(@{const_name HOL.eq}, Type ("fun", [T,_])) $ x $ y) =
 | 
| 16836 | 210 | let val tx = tm x and ty = tm y | 
| 211 | in if pos orelse T = HOLogic.realT then | |
| 212 |                        Buildin("=", [tx, ty])
 | |
| 213 | else if is_intnat T then | |
| 214 |                        Buildin("AND",
 | |
| 215 |                                     [Buildin("<", [tx, suc ty]),
 | |
| 216 |                                      Buildin("<", [ty, suc tx])])
 | |
| 217 | else fail t | |
| 218 | end | |
| 219 | (*inequalities: possible types are nat, int, real*) | |
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 220 |       | fm pos (t as Const(@{const_name Orderings.less},  Type ("fun", [T,_])) $ x $ y) =
 | 
| 16836 | 221 | if not pos orelse T = HOLogic.realT then | 
| 222 |                 Buildin("<", [tm x, tm y])
 | |
| 223 | else if is_intnat T then | |
| 224 |                 Buildin("<=", [suc (tm x), tm y])
 | |
| 225 | else fail t | |
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 226 |       | fm pos (t as Const(@{const_name Orderings.less_eq},  Type ("fun", [T,_])) $ x $ y) =
 | 
| 16836 | 227 | if pos orelse T = HOLogic.realT then | 
| 228 |                 Buildin("<=", [tm x, tm y])
 | |
| 229 | else if is_intnat T then | |
| 230 |                 Buildin("<", [tm x, suc (tm y)])
 | |
| 231 | else fail t | |
| 12869 | 232 | | fm pos t = var(t,[]); | 
| 233 | (*entry point, and translation of a meta-formula*) | |
| 38558 | 234 |       fun mt pos ((c as Const(@{const_name Trueprop}, _)) $ p) = fm pos (iff_tag p)
 | 
| 16836 | 235 |         | mt pos ((c as Const("==>", _)) $ p $ q) =
 | 
| 236 |             Buildin("=>", [mt (not pos) p, mt pos q])
 | |
| 237 | | mt pos t = fm pos (iff_tag t) (*it might be a formula*) | |
| 12869 | 238 | |
| 239 | val body_e = mt pos body (*evaluate now to assign into !nat_vars*) | |
| 16836 | 240 | in | 
| 21078 | 241 | fold_rev add_nat_var (!nat_vars) body_e | 
| 12869 | 242 | end; | 
| 243 | ||
| 244 | ||
| 28290 | 245 | (*The oracle proves the given formula, if possible*) | 
| 246 | fun oracle ct = | |
| 247 | let | |
| 248 | val thy = Thm.theory_of_cterm ct; | |
| 249 | val t = Thm.term_of ct; | |
| 250 | val _ = | |
| 251 |         if ! trace then tracing ("SVC oracle: problem is\n" ^ Syntax.string_of_term_global thy t)
 | |
| 252 | else (); | |
| 253 | in if valid (expr_of false t) then ct else fail t end; | |
| 12869 | 254 | |
| 255 | end; |