src/HOL/ex/Fundefs.thy
changeset 58310 91ea607a34d8
parent 58249 180f1b3508ed
child 58770 ae5e9b4f8daf
equal deleted inserted replaced
58309:a09ec6daaa19 58310:91ea607a34d8
   370 where
   370 where
   371   "f3 x = finite x"
   371   "f3 x = finite x"
   372 
   372 
   373 
   373 
   374 (* Simple Higher-Order Recursion *)
   374 (* Simple Higher-Order Recursion *)
   375 datatype_new 'a tree = 
   375 datatype 'a tree = 
   376   Leaf 'a 
   376   Leaf 'a 
   377   | Branch "'a tree list"
   377   | Branch "'a tree list"
   378 
   378 
   379 fun treemap :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a tree \<Rightarrow> 'a tree"
   379 fun treemap :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a tree \<Rightarrow> 'a tree"
   380 where
   380 where
   421 | "diag True True True = 4"
   421 | "diag True True True = 4"
   422 | "diag False False False = 5"
   422 | "diag False False False = 5"
   423 
   423 
   424 
   424 
   425 (* Many equations (quadratic blowup) *)
   425 (* Many equations (quadratic blowup) *)
   426 datatype_new DT = 
   426 datatype DT = 
   427   A | B | C | D | E | F | G | H | I | J | K | L | M | N | P
   427   A | B | C | D | E | F | G | H | I | J | K | L | M | N | P
   428 | Q | R | S | T | U | V
   428 | Q | R | S | T | U | V
   429 
   429 
   430 fun big :: "DT \<Rightarrow> nat"
   430 fun big :: "DT \<Rightarrow> nat"
   431 where
   431 where