src/ZF/Coind/BCR.thy
author wenzelm
Sat, 27 Oct 2001 00:09:59 +0200
changeset 11963 a6608d44a46b
parent 11318 6536fb8c9fc6
child 12595 0480d02221b8
permissions -rw-r--r--
impose hyps on initial goal configuration (prevents res_inst_tac problems);

(*  Title:      ZF/Coind/BCR.thy
    ID:         $Id$
    Author:     Jacob Frost, Cambridge University Computer Laboratory
    Copyright   1995  University of Cambridge
*)

BCR = Values + Types +

(*Basic correspondence relation -- not completely specified, as it is a
  parameter of the proof.  A concrete version could be defined inductively.*)

consts
  isof :: [i,i] => o
rules
  isof_app "[| isof(c1,t_fun(t1,t2)); isof(c2,t1) |] ==> isof(c_app(c1,c2),t2)"

(*Its extension to environments*)

consts
  isofenv :: [i,i] => o
defs
  isofenv_def "isofenv(ve,te) ==                
   ve_dom(ve) = te_dom(te) &            
   ( \\<forall>x \\<in> ve_dom(ve).                          
     \\<exists>c \\<in> Const. ve_app(ve,x) = v_const(c) & isof(c,te_app(te,x)))"
  
end