src/HOL/BNF/Tools/bnf_fp_rec_sugar_util.ML
author blanchet
Mon, 04 Nov 2013 16:53:43 +0100
changeset 55698 8fdb4dc08ed1
parent 55695 a596292be9a8
child 55751 bc24e1ccfd35
permissions -rw-r--r--
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
blanchet@54440
     1
(*  Title:      HOL/BNF/Tools/bnf_fp_rec_sugar_util.ML
blanchet@54440
     2
    Author:     Lorenz Panny, TU Muenchen
blanchet@54440
     3
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@54440
     4
    Copyright   2013
blanchet@54440
     5
blanchet@54440
     6
Library for recursor and corecursor sugar.
blanchet@54440
     7
*)
blanchet@54440
     8
blanchet@54440
     9
signature BNF_FP_REC_SUGAR_UTIL =
blanchet@54440
    10
sig
blanchet@55698
    11
  val indexed: 'a list -> int -> int list * int
blanchet@55698
    12
  val indexedd: 'a list list -> int -> int list list * int
blanchet@55698
    13
  val indexeddd: ''a list list list -> int -> int list list list * int
blanchet@55698
    14
  val indexedddd: 'a list list list list -> int -> int list list list list * int
blanchet@55698
    15
  val find_index_eq: ''a list -> ''a -> int
blanchet@55698
    16
  val finds: ('a * 'b -> bool) -> 'a list -> 'b list -> ('a * 'b list) list * 'b list
blanchet@54440
    17
blanchet@55698
    18
  val drop_All: term -> term
blanchet@54440
    19
blanchet@55698
    20
  val mk_partial_compN: int -> typ -> term -> term
blanchet@55698
    21
  val mk_partial_comp: typ -> typ -> term -> term
blanchet@55698
    22
  val mk_compN: int -> typ list -> term * term -> term
blanchet@55698
    23
  val mk_comp: typ list -> term * term -> term
blanchet@54440
    24
blanchet@55698
    25
  val get_indices: ((binding * typ) * 'a) list -> term -> int list
blanchet@54440
    26
end;
blanchet@54440
    27
blanchet@54440
    28
structure BNF_FP_Rec_Sugar_Util : BNF_FP_REC_SUGAR_UTIL =
blanchet@54440
    29
struct
blanchet@54440
    30
blanchet@54440
    31
fun indexed xs h = let val h' = h + length xs in (h upto h' - 1, h') end;
blanchet@54440
    32
fun indexedd xss = fold_map indexed xss;
blanchet@54440
    33
fun indexeddd xsss = fold_map indexedd xsss;
blanchet@54440
    34
fun indexedddd xssss = fold_map indexeddd xssss;
blanchet@54440
    35
blanchet@54440
    36
fun find_index_eq hs h = find_index (curry (op =) h) hs;
blanchet@54440
    37
blanchet@55698
    38
fun finds eq = fold_map (fn x => List.partition (curry eq x) #>> pair x);
blanchet@54613
    39
blanchet@55698
    40
fun drop_All t =
blanchet@55698
    41
  subst_bounds (strip_qnt_vars @{const_name all} t |> map Free |> rev,
blanchet@55698
    42
    strip_qnt_body @{const_name all} t);
blanchet@54440
    43
blanchet@55698
    44
fun mk_partial_comp gT fT g =
blanchet@55698
    45
  let val T = domain_type fT --> range_type gT in
blanchet@55698
    46
    Const (@{const_name Fun.comp}, gT --> fT --> T) $ g
blanchet@54440
    47
  end;
blanchet@54440
    48
blanchet@55698
    49
fun mk_partial_compN 0 _ g = g
blanchet@55698
    50
  | mk_partial_compN n fT g =
blanchet@55698
    51
    let val g' = mk_partial_compN (n - 1) (range_type fT) g in
blanchet@55698
    52
      mk_partial_comp (fastype_of g') fT g'
blanchet@55698
    53
    end;
blanchet@55586
    54
blanchet@55698
    55
fun mk_compN n bound_Ts (g, f) =
blanchet@55698
    56
  let val typof = curry fastype_of1 bound_Ts in
blanchet@55698
    57
    mk_partial_compN n (typof f) g $ f
blanchet@54440
    58
  end;
blanchet@54440
    59
blanchet@55698
    60
val mk_comp = mk_compN 1;
blanchet@55698
    61
blanchet@55698
    62
fun get_indices fixes t = map (fst #>> Binding.name_of #> Free) fixes
blanchet@55698
    63
  |> map_index (fn (i, v) => if exists_subterm (equal v) t then SOME i else NONE)
blanchet@55698
    64
  |> map_filter I;
blanchet@55698
    65
blanchet@54440
    66
end;