src/HOL/Tools/Function/sum_tree.ML
author wenzelm
Wed, 06 Dec 2017 20:43:09 +0100
changeset 67149 e61557884799
parent 55968 94242fa87638
child 69593 3dda49e08b9d
permissions -rw-r--r--
prefer control symbol antiquotations;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31775
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 29183
diff changeset
     1
(*  Title:      HOL/Tools/Function/sum_tree.ML
25555
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
     2
    Author:     Alexander Krauss, TU Muenchen
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
     3
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
     4
Some common tools for working with sum types in balanced tree form.
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
     5
*)
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
     6
49965
ee25a04fa06a proper signatures;
wenzelm
parents: 37678
diff changeset
     7
signature SUM_TREE =
ee25a04fa06a proper signatures;
wenzelm
parents: 37678
diff changeset
     8
sig
ee25a04fa06a proper signatures;
wenzelm
parents: 37678
diff changeset
     9
  val sumcase_split_ss: simpset
ee25a04fa06a proper signatures;
wenzelm
parents: 37678
diff changeset
    10
  val access_top_down: {init: 'a, left: 'a -> 'a, right: 'a -> 'a} -> int -> int -> 'a
ee25a04fa06a proper signatures;
wenzelm
parents: 37678
diff changeset
    11
  val mk_sumT: typ -> typ -> typ
ee25a04fa06a proper signatures;
wenzelm
parents: 37678
diff changeset
    12
  val mk_sumcase: typ -> typ -> typ -> term -> term -> term
ee25a04fa06a proper signatures;
wenzelm
parents: 37678
diff changeset
    13
  val App: term -> term -> term
ee25a04fa06a proper signatures;
wenzelm
parents: 37678
diff changeset
    14
  val mk_inj: typ -> int -> int -> term -> term
ee25a04fa06a proper signatures;
wenzelm
parents: 37678
diff changeset
    15
  val mk_proj: typ -> int -> int -> term -> term
ee25a04fa06a proper signatures;
wenzelm
parents: 37678
diff changeset
    16
  val mk_sumcases: typ -> term list -> term
ee25a04fa06a proper signatures;
wenzelm
parents: 37678
diff changeset
    17
end
ee25a04fa06a proper signatures;
wenzelm
parents: 37678
diff changeset
    18
55968
blanchet
parents: 55642
diff changeset
    19
structure Sum_Tree: SUM_TREE =
25555
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    20
struct
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    21
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    22
(* Theory dependencies *)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33961
diff changeset
    23
val sumcase_split_ss =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49965
diff changeset
    24
  simpset_of (put_simpset HOL_basic_ss @{context}
55642
63beb38e9258 adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents: 55414
diff changeset
    25
    addsimps (@{thm Product_Type.split} :: @{thms sum.case}))
25555
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    26
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    27
(* top-down access in balanced tree *)
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    28
fun access_top_down {left, right, init} len i =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33961
diff changeset
    29
  Balanced_Tree.access
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33961
diff changeset
    30
    {left = (fn f => f o left), right = (fn f => f o right), init = I} len i init
25555
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
(* Sum types *)
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 55968
diff changeset
    33
fun mk_sumT LT RT = Type (\<^type_name>\<open>Sum_Type.sum\<close>, [LT, RT])
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33961
diff changeset
    34
fun mk_sumcase TL TR T l r =
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 55968
diff changeset
    35
  Const (\<^const_name>\<open>sum.case_sum\<close>, (TL --> T) --> (TR --> T) --> mk_sumT TL TR --> T) $ l $ r
25555
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    36
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    37
val App = curry op $
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    38
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33961
diff changeset
    39
fun mk_inj ST n i =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33961
diff changeset
    40
  access_top_down
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33961
diff changeset
    41
  { init = (ST, I : term -> term),
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 55968
diff changeset
    42
    left = (fn (T as Type (\<^type_name>\<open>Sum_Type.sum\<close>, [LT, RT]), inj) =>
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 55968
diff changeset
    43
      (LT, inj o App (Const (\<^const_name>\<open>Inl\<close>, LT --> T)))),
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 55968
diff changeset
    44
    right =(fn (T as Type (\<^type_name>\<open>Sum_Type.sum\<close>, [LT, RT]), inj) =>
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 55968
diff changeset
    45
      (RT, inj o App (Const (\<^const_name>\<open>Inr\<close>, RT --> T))))} n i
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33961
diff changeset
    46
  |> snd
25555
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    47
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33961
diff changeset
    48
fun mk_proj ST n i =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33961
diff changeset
    49
  access_top_down
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33961
diff changeset
    50
  { init = (ST, I : term -> term),
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 55968
diff changeset
    51
    left = (fn (T as Type (\<^type_name>\<open>Sum_Type.sum\<close>, [LT, RT]), proj) =>
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 55968
diff changeset
    52
      (LT, App (Const (\<^const_name>\<open>Sum_Type.projl\<close>, T --> LT)) o proj)),
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 55968
diff changeset
    53
    right =(fn (T as Type (\<^type_name>\<open>Sum_Type.sum\<close>, [LT, RT]), proj) =>
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 55968
diff changeset
    54
      (RT, App (Const (\<^const_name>\<open>Sum_Type.projr\<close>, T --> RT)) o proj))} n i
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33961
diff changeset
    55
  |> snd
25555
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    56
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    57
fun mk_sumcases T fs =
49965
ee25a04fa06a proper signatures;
wenzelm
parents: 37678
diff changeset
    58
  Balanced_Tree.make (fn ((f, fT), (g, gT)) => (mk_sumcase fT gT T f g, mk_sumT fT gT))
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33961
diff changeset
    59
    (map (fn f => (f, domain_type (fastype_of f))) fs)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33961
diff changeset
    60
  |> fst
25555
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    61
224a40e39457 factored out handling of sum types again
krauss
parents:
diff changeset
    62
end