blanchet@58309: (* Title: HOL/Datatype_Examples/Misc_Primcorec.thy blanchet@54193: Author: Jasmin Blanchette, TU Muenchen blanchet@54193: Copyright 2013 blanchet@54193: blanchet@54193: Miscellaneous primitive corecursive function definitions. blanchet@54193: *) blanchet@54193: wenzelm@58889: section {* Miscellaneous Primitive Corecursive Function Definitions *} blanchet@54193: blanchet@54193: theory Misc_Primcorec blanchet@54193: imports Misc_Codatatype blanchet@54193: begin blanchet@54193: blanchet@54193: primcorec simple_of_bools :: "bool \ bool \ simple" where blanchet@54193: "simple_of_bools b b' = (if b then if b' then X1 else X2 else if b' then X3 else X4)" blanchet@54193: blanchet@54193: primcorec simple'_of_bools :: "bool \ bool \ simple'" where blanchet@54193: "simple'_of_bools b b' = blanchet@54193: (if b then if b' then X1' () else X2' () else if b' then X3' () else X4' ())" blanchet@54193: blanchet@54193: primcorec inc_simple'' :: "nat \ simple'' \ simple''" where blanchet@54193: "inc_simple'' k s = (case s of X1'' n i \ X1'' (n + k) (i + int k) | X2'' \ X2'')" blanchet@54193: blanchet@54193: primcorec sinterleave :: "'a stream \ 'a stream \ 'a stream" where blanchet@54193: "sinterleave s s' = Stream (shd s) (sinterleave s' (stl s))" blanchet@54193: blanchet@54193: primcorec myapp :: "'a mylist \ 'a mylist \ 'a mylist" where blanchet@54193: "myapp xs ys = blanchet@54193: (if xs = MyNil then ys blanchet@54193: else if ys = MyNil then xs blanchet@54193: else MyCons (myhd xs) (myapp (mytl xs) ys))" blanchet@54193: blanchet@58231: primcorec shuffle_sp :: "('a \ ord, 'b \ ord, 'c, 'd) some_passive \ ('d, 'a, 'b, 'c) some_passive" where blanchet@54193: "shuffle_sp sp = blanchet@54193: (case sp of blanchet@54193: SP1 sp' \ SP1 (shuffle_sp sp') blanchet@54193: | SP2 a \ SP3 a blanchet@54193: | SP3 b \ SP4 b blanchet@54193: | SP4 c \ SP5 c blanchet@54193: | SP5 d \ SP2 d)" blanchet@54193: blanchet@54193: primcorec rename_lam :: "(string \ string) \ lambda \ lambda" where blanchet@54193: "rename_lam f l = blanchet@54193: (case l of blanchet@54193: Var s \ Var (f s) blanchet@54193: | App l l' \ App (rename_lam f l) (rename_lam f l') blanchet@54193: | Abs s l \ Abs (f s) (rename_lam f l) blanchet@55932: | Let SL l \ Let (fimage (map_prod f (rename_lam f)) SL) (rename_lam f l))" blanchet@54193: blanchet@54193: primcorec blanchet@54193: j1_sum :: "('a\{zero,one,plus}) \ 'a J1" and blanchet@54193: j2_sum :: "'a \ 'a J2" blanchet@54193: where blanchet@54193: "n = 0 \ is_J11 (j1_sum n)" | blanchet@54193: "un_J111 (j1_sum _) = 0" | blanchet@54193: "un_J112 (j1_sum _) = j1_sum 0" | blanchet@54193: "un_J121 (j1_sum n) = n + 1" | blanchet@54193: "un_J122 (j1_sum n) = j2_sum (n + 1)" | blanchet@57090: "n = 0 \ j2_sum n = J21" | blanchet@54193: "un_J221 (j2_sum n) = j1_sum (n + 1)" | blanchet@54193: "un_J222 (j2_sum n) = j2_sum (n + 1)" blanchet@54193: blanchet@54193: primcorec forest_of_mylist :: "'a tree mylist \ 'a forest" where blanchet@54193: "forest_of_mylist ts = blanchet@54193: (case ts of blanchet@54193: MyNil \ FNil blanchet@54193: | MyCons t ts \ FCons t (forest_of_mylist ts))" blanchet@54193: blanchet@54193: primcorec mylist_of_forest :: "'a forest \ 'a tree mylist" where blanchet@54193: "mylist_of_forest f = blanchet@54193: (case f of blanchet@54193: FNil \ MyNil blanchet@54193: | FCons t ts \ MyCons t (mylist_of_forest ts))" blanchet@54193: blanchet@54193: primcorec semi_stream :: "'a stream \ 'a stream" where blanchet@54193: "semi_stream s = Stream (shd s) (semi_stream (stl (stl s)))" blanchet@54193: blanchet@54193: primcorec blanchet@54193: tree'_of_stream :: "'a stream \ 'a tree'" and blanchet@54193: branch_of_stream :: "'a stream \ 'a branch" blanchet@54193: where blanchet@54193: "tree'_of_stream s = blanchet@54193: TNode' (branch_of_stream (semi_stream s)) (branch_of_stream (semi_stream (stl s)))" | blanchet@54193: "branch_of_stream s = (case s of Stream h t \ Branch h (tree'_of_stream t))" blanchet@54193: blanchet@54193: primcorec blanchet@55484: id_tree :: "'a bin_rose_tree \ 'a bin_rose_tree" and blanchet@55484: id_trees1 :: "'a bin_rose_tree mylist \ 'a bin_rose_tree mylist" and blanchet@55484: id_trees2 :: "'a bin_rose_tree mylist \ 'a bin_rose_tree mylist" blanchet@55484: where blanchet@55852: "id_tree t = (case t of BRTree a ts ts' \ BRTree a (id_trees1 ts) (id_trees2 ts'))" | blanchet@55484: "id_trees1 ts = (case ts of blanchet@55852: MyNil \ MyNil blanchet@55852: | MyCons t ts \ MyCons (id_tree t) (id_trees1 ts))" | blanchet@55484: "id_trees2 ts = (case ts of blanchet@55852: MyNil \ MyNil blanchet@55852: | MyCons t ts \ MyCons (id_tree t) (id_trees2 ts))" blanchet@55484: blanchet@55484: primcorec blanchet@55774: trunc_tree :: "'a bin_rose_tree \ 'a bin_rose_tree" and blanchet@55774: trunc_trees1 :: "'a bin_rose_tree mylist \ 'a bin_rose_tree mylist" and blanchet@55774: trunc_trees2 :: "'a bin_rose_tree mylist \ 'a bin_rose_tree mylist" blanchet@55774: where blanchet@55852: "trunc_tree t = (case t of BRTree a ts ts' \ BRTree a (trunc_trees1 ts) (trunc_trees2 ts'))" | blanchet@55774: "trunc_trees1 ts = (case ts of blanchet@55852: MyNil \ MyNil blanchet@55852: | MyCons t _ \ MyCons (trunc_tree t) MyNil)" | blanchet@55774: "trunc_trees2 ts = (case ts of blanchet@55852: MyNil \ MyNil blanchet@55852: | MyCons t ts \ MyCons (trunc_tree t) MyNil)" blanchet@55774: blanchet@55774: primcorec blanchet@54193: freeze_exp :: "('b \ 'a) \ ('a, 'b) exp \ ('a, 'b) exp" and blanchet@54193: freeze_trm :: "('b \ 'a) \ ('a, 'b) trm \ ('a, 'b) trm" and blanchet@54193: freeze_factor :: "('b \ 'a) \ ('a, 'b) factor \ ('a, 'b) factor" blanchet@54193: where blanchet@54193: "freeze_exp g e = blanchet@54193: (case e of blanchet@54193: Term t \ Term (freeze_trm g t) blanchet@54193: | Sum t e \ Sum (freeze_trm g t) (freeze_exp g e))" | blanchet@54193: "freeze_trm g t = blanchet@54193: (case t of blanchet@54193: Factor f \ Factor (freeze_factor g f) blanchet@54193: | Prod f t \ Prod (freeze_factor g f) (freeze_trm g t))" | blanchet@54193: "freeze_factor g f = blanchet@54193: (case f of blanchet@54193: C a \ C a blanchet@54193: | V b \ C (g b) blanchet@54193: | Paren e \ Paren (freeze_exp g e))" blanchet@54193: blanchet@54193: primcorec poly_unity :: "'a poly_unit" where blanchet@54193: "poly_unity = U (\_. poly_unity)" blanchet@54193: blanchet@54193: primcorec build_cps :: "('a \ 'a) \ ('a \ bool stream) \ 'a \ bool stream \ 'a cps" where blanchet@54193: "shd b \ build_cps f g a b = CPS1 a" | blanchet@54193: "_ \ build_cps f g a b = CPS2 (\a. build_cps f g (f a) (g a))" blanchet@54193: blanchet@54193: end