src/HOL/List.thy
changeset 57123 b5324647e0f1
parent 57091 1fa9c19ba2c9
child 57198 159e1b043495
equal deleted inserted replaced
57122:5f69b8c3af5a 57123:b5324647e0f1
     6 
     6 
     7 theory List
     7 theory List
     8 imports Presburger Code_Numeral Quotient Lifting_Set Lifting_Option Lifting_Product
     8 imports Presburger Code_Numeral Quotient Lifting_Set Lifting_Option Lifting_Product
     9 begin
     9 begin
    10 
    10 
    11 datatype_new (set: 'a) list (map: map rel: list_all2) =
    11 datatype_new (set: 'a) list  (map: map rel: list_all2) =
    12     Nil (defaults tl: "[]")  ("[]")
    12     Nil (defaults tl: "[]")  ("[]")
    13   | Cons (hd: 'a) (tl: "'a list")  (infixr "#" 65)
    13   | Cons (hd: 'a) (tl: "'a list")  (infixr "#" 65)
       
    14 
    14 datatype_compat list
    15 datatype_compat list
    15 
    16 
    16 lemma [case_names Nil Cons, cases type: list]:
    17 lemma [case_names Nil Cons, cases type: list]:
    17   -- {* for backward compatibility -- names of variables differ *}
    18   -- {* for backward compatibility -- names of variables differ *}
    18   "(y = [] \<Longrightarrow> P) \<Longrightarrow> (\<And>a list. y = a # list \<Longrightarrow> P) \<Longrightarrow> P"
    19   "(y = [] \<Longrightarrow> P) \<Longrightarrow> (\<And>a list. y = a # list \<Longrightarrow> P) \<Longrightarrow> P"