src/HOL/Library/Finite_Map.thy
changeset 67780 7655e6369c9f
parent 67485 89f5d876a656
child 68249 949d93804740
equal deleted inserted replaced
67779:fd2558014196 67780:7655e6369c9f
     4 
     4 
     5 section \<open>Type of finite maps defined as a subtype of maps\<close>
     5 section \<open>Type of finite maps defined as a subtype of maps\<close>
     6 
     6 
     7 theory Finite_Map
     7 theory Finite_Map
     8   imports FSet AList
     8   imports FSet AList
       
     9   abbrevs "(=" = "\<subseteq>\<^sub>f"
     9 begin
    10 begin
    10 
    11 
    11 subsection \<open>Auxiliary constants and lemmas over @{type map}\<close>
    12 subsection \<open>Auxiliary constants and lemmas over @{type map}\<close>
    12 
    13 
    13 context includes lifting_syntax begin
    14 context includes lifting_syntax begin