src/HOL/HOLCF/Universal.thy
changeset 66453 cc19f7ca2ed6
parent 65552 f533820e7248
child 67399 eab6ce8368fa
equal deleted inserted replaced
66452:450cefec7c11 66453:cc19f7ca2ed6
     3 *)
     3 *)
     4 
     4 
     5 section \<open>A universal bifinite domain\<close>
     5 section \<open>A universal bifinite domain\<close>
     6 
     6 
     7 theory Universal
     7 theory Universal
     8 imports Bifinite Completion "~~/src/HOL/Library/Nat_Bijection"
     8 imports Bifinite Completion "HOL-Library.Nat_Bijection"
     9 begin
     9 begin
    10 
    10 
    11 no_notation binomial  (infixl "choose" 65)
    11 no_notation binomial  (infixl "choose" 65)
    12 
    12 
    13 subsection \<open>Basis for universal domain\<close>
    13 subsection \<open>Basis for universal domain\<close>