src/HOL/UNITY/Comp/Counterc.thy
changeset 63146 f1ecba0272f9
parent 58889 5b7a9633cfa8
equal deleted inserted replaced
63145:703edebd1d92 63146:f1ecba0272f9
     9 Examples of Program Composition Illustrating the Use of Universal Properties
     9 Examples of Program Composition Illustrating the Use of Universal Properties
    10    In J. Rolim (editor), Parallel and Distributed Processing,
    10    In J. Rolim (editor), Parallel and Distributed Processing,
    11    Spriner LNCS 1586 (1999), pages 1215-1227.
    11    Spriner LNCS 1586 (1999), pages 1215-1227.
    12 *)
    12 *)
    13 
    13 
    14 section{*A Family of Similar Counters: Version with Compatibility*}
    14 section\<open>A Family of Similar Counters: Version with Compatibility\<close>
    15 
    15 
    16 theory Counterc imports "../UNITY_Main" begin
    16 theory Counterc imports "../UNITY_Main" begin
    17 
    17 
    18 typedecl state
    18 typedecl state
    19 
    19