src/HOL/Tools/function_package/sum_tree.ML
author wenzelm
Wed, 31 Dec 2008 18:53:16 +0100
changeset 29270 0eade173f77e
parent 29183 f1648e009dc1
permissions -rw-r--r--
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
25555
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
     1
(*  Title:      HOL/Tools/function_package/sum_tree.ML
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
     2
    ID:         $Id$
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
     3
    Author:     Alexander Krauss, TU Muenchen
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
     4
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
     5
Some common tools for working with sum types in balanced tree form.
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
     6
*)
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
     7
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
     8
structure SumTree =
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
     9
struct
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    10
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    11
(* Theory dependencies *)
29183
f1648e009dc1 removed duplicate sum_case used only by function package;
krauss
parents: 25555
diff changeset
    12
val proj_in_rules = [@{thm "Datatype.Projl_Inl"}, @{thm "Datatype.Projr_Inr"}]
f1648e009dc1 removed duplicate sum_case used only by function package;
krauss
parents: 25555
diff changeset
    13
val sumcase_split_ss = HOL_basic_ss addsimps (@{thm "Product_Type.split"} :: @{thms "sum.cases"})
25555
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    14
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    15
(* top-down access in balanced tree *)
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    16
fun access_top_down {left, right, init} len i =
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    17
    BalancedTree.access {left = (fn f => f o left), right = (fn f => f o right), init = I} len i init
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    18
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    19
(* Sum types *)
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    20
fun mk_sumT LT RT = Type ("+", [LT, RT])
29183
f1648e009dc1 removed duplicate sum_case used only by function package;
krauss
parents: 25555
diff changeset
    21
fun mk_sumcase TL TR T l r = Const (@{const_name "sum.sum_case"}, (TL --> T) --> (TR --> T) --> mk_sumT TL TR --> T) $ l $ r
25555
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    22
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    23
val App = curry op $
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    24
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    25
fun mk_inj ST n i = 
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    26
    access_top_down 
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    27
    { init = (ST, I : term -> term),
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    28
      left = (fn (T as Type ("+", [LT, RT]), inj) => (LT, inj o App (Const (@{const_name "Inl"}, LT --> T)))),
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    29
      right =(fn (T as Type ("+", [LT, RT]), inj) => (RT, inj o App (Const (@{const_name "Inr"}, RT --> T))))} n i 
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    30
    |> snd
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    31
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    32
fun mk_proj ST n i = 
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    33
    access_top_down 
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    34
    { init = (ST, I : term -> term),
29183
f1648e009dc1 removed duplicate sum_case used only by function package;
krauss
parents: 25555
diff changeset
    35
      left = (fn (T as Type ("+", [LT, RT]), proj) => (LT, App (Const (@{const_name "Datatype.Projl"}, T --> LT)) o proj)),
f1648e009dc1 removed duplicate sum_case used only by function package;
krauss
parents: 25555
diff changeset
    36
      right =(fn (T as Type ("+", [LT, RT]), proj) => (RT, App (Const (@{const_name "Datatype.Projr"}, T --> RT)) o proj))} n i
25555
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    37
    |> snd
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    38
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    39
fun mk_sumcases T fs =
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    40
      BalancedTree.make (fn ((f, fT), (g, gT)) => (mk_sumcase fT gT T f g, mk_sumT fT gT)) 
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    41
                        (map (fn f => (f, domain_type (fastype_of f))) fs)
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    42
      |> fst
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    43
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    44
end