src/HOL/HOLCF/IOA/ABP/Env.thy
changeset 62009 ecb5212d5885
parent 62008 cbedaddc9351
child 66453 cc19f7ca2ed6
equal deleted inserted replaced
62008:cbedaddc9351 62009:ecb5212d5885
     3 *)
     3 *)
     4 
     4 
     5 section \<open>The environment\<close>
     5 section \<open>The environment\<close>
     6 
     6 
     7 theory Env
     7 theory Env
     8 imports "~~/src/HOL/HOLCF/IOA/IOA" Action
     8 imports "../IOA" Action
     9 begin
     9 begin
    10 
    10 
    11 type_synonym
    11 type_synonym
    12   'm env_state = bool   \<comment> \<open>give next bit to system\<close>
    12   'm env_state = bool   \<comment> \<open>give next bit to system\<close>
    13 
    13