src/ZF/UNITY/State.thy
changeset 60770 240563fbf41d
parent 58871 c399ae4b836f
child 61980 6b780867d426
equal deleted inserted replaced
60769:cf7f3465eaf1 60770:240563fbf41d
     6  - variables are typed.
     6  - variables are typed.
     7  - the state space is uniform, common to all defined programs.
     7  - the state space is uniform, common to all defined programs.
     8  - variables can be quantified over.
     8  - variables can be quantified over.
     9 *)
     9 *)
    10 
    10 
    11 section{*UNITY Program States*}
    11 section\<open>UNITY Program States\<close>
    12 
    12 
    13 theory State imports Main begin
    13 theory State imports Main begin
    14 
    14 
    15 consts var :: i
    15 consts var :: i
    16 datatype var = Var("i \<in> list(nat)")
    16 datatype var = Var("i \<in> list(nat)")