src/HOL/TLA/Memory/MemoryParameters.ML
author nipkow
Mon, 06 Aug 2001 13:43:24 +0200
changeset 11464 ddea204de5bc
parent 9517 f58863b1406a
child 17309 c43ed29bd197
permissions -rw-r--r--
turned translation for 1::nat into def. introduced 1' and replaced most occurrences of 1 by 1'.

(* 
    File:        MemoryParameters.ML
    Author:      Stephan Merz
    Copyright:   1997 University of Munich

    RPC-Memory example: memory parameters (ML file)
*)

Addsimps ([BadArgNoMemVal,MemFailNoMemVal,InitValMemVal,NotAResultNotVal,
                  NotAResultNotOK, NotAResultNotBA, NotAResultNotMF]
               @ (map (fn x => x RS not_sym) 
                      [NotAResultNotOK, NotAResultNotBA, NotAResultNotMF]));

val prems = goal thy "[| x : MemVal; (x ~= NotAResult ==> P) |] ==> P";
by (resolve_tac prems 1);
by (cut_facts_tac (NotAResultNotVal::prems) 1);
by (Force_tac 1);
qed "MemValNotAResultE";