src/HOL/Library/case_converter.ML
author wenzelm
Fri, 04 Jan 2019 23:22:53 +0100
changeset 69593 3dda49e08b9d
parent 69568 de09a7261120
permissions -rw-r--r--
isabelle update -u control_cartouches;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
68155
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
     1
(* Author: Pascal Stoop, ETH Zurich
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
     2
   Author: Andreas Lochbihler, Digital Asset *)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
     3
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
     4
signature CASE_CONVERTER =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
     5
sig
69568
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
     6
  type elimination_strategy
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
     7
  val to_case: Proof.context -> elimination_strategy -> (string * typ -> int) ->
68155
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
     8
    thm list -> thm list option
69568
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
     9
  val replace_by_type: (Proof.context -> string * string -> bool) -> elimination_strategy
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    10
  val keep_constructor_context: elimination_strategy
68155
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    11
end;
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    12
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    13
structure Case_Converter : CASE_CONVERTER =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    14
struct
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    15
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    16
fun lookup_remove _ _ [] = (NONE, [])
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    17
  | lookup_remove eq k ((k', v) :: kvs) =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    18
    if eq (k, k') then (SOME (k', v), kvs)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    19
    else apsnd (cons (k', v)) (lookup_remove eq k kvs)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    20
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    21
fun mk_abort msg t =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    22
  let 
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    23
    val T = fastype_of t
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69568
diff changeset
    24
    val abort = Const (\<^const_name>\<open>missing_pattern_match\<close>, HOLogic.literalT --> (HOLogic.unitT --> T) --> T)
68155
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    25
  in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    26
    abort $ HOLogic.mk_literal msg $ absdummy HOLogic.unitT t
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    27
  end
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    28
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    29
(* fold_term : (string * typ -> 'a) ->
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    30
               (string * typ -> 'a) ->
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    31
               (indexname * typ -> 'a) ->
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    32
               (int -> 'a) ->
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    33
               (string * typ * 'a -> 'a) ->
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    34
               ('a * 'a -> 'a) ->
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    35
               term ->
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    36
               'a *)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    37
fun fold_term const_fun free_fun var_fun bound_fun abs_fun dollar_fun term =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    38
  let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    39
    fun go x = case x of
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    40
      Const (s, T) => const_fun (s, T)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    41
      | Free (s, T) => free_fun (s, T)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    42
      | Var (i, T) => var_fun (i, T)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    43
      | Bound n => bound_fun n
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    44
      | Abs (s, T, term) => abs_fun (s, T, go term)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    45
      | term1 $ term2 => dollar_fun (go term1, go term2)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    46
  in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    47
    go term
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    48
  end;
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    49
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    50
datatype term_coordinate = End of typ
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    51
  | Coordinate of (string * (int * term_coordinate)) list;
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    52
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    53
fun term_coordinate_merge (End T) _ = End T
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    54
  | term_coordinate_merge _ (End T) = End T
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    55
  | term_coordinate_merge (Coordinate xs) (Coordinate ys) =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    56
  let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    57
    fun merge_consts xs [] = xs
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    58
      | merge_consts xs ((s1, (n, y)) :: ys) = 
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    59
        case List.partition (fn (s2, (m, _)) => s1 = s2 andalso n = m) xs of
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    60
            ([], xs') => (s1, (n, y)) :: (merge_consts xs' ys)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    61
          | ((_, (_, x)) :: _, xs') => (s1, (n, term_coordinate_merge x y)) :: (merge_consts xs' ys)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    62
  in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    63
    Coordinate (merge_consts xs ys)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    64
  end;
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    65
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    66
fun coordinates_to_list (End x) = [(x, [])]
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    67
  | coordinates_to_list (Coordinate xs) = 
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    68
  let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    69
    fun f (s, (n, xss)) = map (fn (T, xs) => (T, (s, n) :: xs)) (coordinates_to_list xss)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    70
  in flat (map f xs) end;
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
    71
69568
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    72
type elimination_strategy = Proof.context -> term list -> term_coordinate list
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    73
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    74
fun replace_by_type replace_ctr ctxt pats =
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    75
  let
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    76
    fun term_to_coordinates P term = 
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    77
      let
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    78
        val (ctr, args) = strip_comb term
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    79
      in
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    80
        case ctr of Const (s, T) =>
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    81
          if P (body_type T |> dest_Type |> fst, s)
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    82
          then SOME (End (body_type T))
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    83
          else
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    84
            let
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    85
              fun f (i, t) = term_to_coordinates P t |> Option.map (pair i)
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    86
              val tcos = map_filter I (map_index f args)
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    87
            in
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    88
              if null tcos then NONE
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    89
              else SOME (Coordinate (map (pair s) tcos))
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    90
            end
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    91
        | _ => NONE
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    92
      end
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    93
    in
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    94
      map_filter (term_to_coordinates (replace_ctr ctxt)) pats
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    95
    end
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    96
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    97
fun keep_constructor_context ctxt pats =
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    98
  let
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
    99
    fun to_coordinates [] = NONE
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   100
      | to_coordinates pats =
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   101
        let
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   102
          val (fs, argss) = map strip_comb pats |> split_list
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   103
          val f = hd fs
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   104
          fun is_single_ctr (Const (name, T)) = 
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   105
              let
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   106
                val tyco = body_type T |> dest_Type |> fst
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   107
                val _ = Ctr_Sugar.ctr_sugar_of ctxt tyco |> the |> #ctrs
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   108
              in
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   109
                case Ctr_Sugar.ctr_sugar_of ctxt tyco of
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   110
                  NONE => error ("Not a free constructor " ^ name ^ " in pattern")
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   111
                | SOME info =>
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   112
                  case #ctrs info of [Const (name', _)] => name = name'
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   113
                    | _ => false
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   114
              end
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   115
            | is_single_ctr _ = false
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   116
        in 
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   117
          if not (is_single_ctr f) andalso forall (fn x => f = x) fs then
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   118
            let
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   119
              val patss = Ctr_Sugar_Util.transpose argss
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   120
              fun recurse (i, pats) = to_coordinates pats |> Option.map (pair i)
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   121
              val coords = map_filter I (map_index recurse patss)
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   122
            in
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   123
              if null coords then NONE
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   124
              else SOME (Coordinate (map (pair (dest_Const f |> fst)) coords))
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   125
            end
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   126
          else SOME (End (body_type (fastype_of f)))
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   127
          end
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   128
    in
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   129
      the_list (to_coordinates pats)
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   130
    end
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   131
68155
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   132
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   133
(* AL: TODO: change from term to const_name *)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   134
fun find_ctr ctr1 xs =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   135
  let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   136
    val const_name = fst o dest_Const
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   137
    fun const_equal (ctr1, ctr2) = const_name ctr1 = const_name ctr2
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   138
  in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   139
    lookup_remove const_equal ctr1 xs
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   140
  end;
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   141
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   142
datatype pattern 
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   143
  = Wildcard
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   144
  | Value
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   145
  | Split of int * (term * pattern) list * pattern;
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   146
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   147
fun pattern_merge Wildcard pat' = pat'
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   148
  | pattern_merge Value _ = Value
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   149
  | pattern_merge (Split (n, xs, pat)) Wildcard =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   150
    Split (n, map (apsnd (fn pat'' => pattern_merge pat'' Wildcard)) xs, pattern_merge pat Wildcard)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   151
  | pattern_merge (Split _) Value = Value
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   152
  | pattern_merge (Split (n, xs, pat)) (Split (m, ys, pat'')) =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   153
    let 
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   154
      fun merge_consts xs [] = map (apsnd (fn pat => pattern_merge pat Wildcard)) xs
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   155
        | merge_consts xs ((ctr, y) :: ys) =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   156
          (case find_ctr ctr xs of
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   157
              (SOME (ctr, x), xs) => (ctr, pattern_merge x y) :: merge_consts xs ys
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   158
            | (NONE, xs) => (ctr, y) :: merge_consts xs ys
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   159
          )
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   160
     in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   161
       Split (if n <= 0 then m else n, merge_consts xs ys, pattern_merge pat pat'')
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   162
     end
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   163
     
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   164
fun pattern_intersect Wildcard _ = Wildcard
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   165
  | pattern_intersect Value pat2 = pat2
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   166
  | pattern_intersect (Split _) Wildcard = Wildcard
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   167
  | pattern_intersect (Split (n, xs', pat1)) Value =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   168
    Split (n,
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   169
      map (apsnd (fn pat1 => pattern_intersect pat1 Value)) xs',
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   170
      pattern_intersect pat1 Value)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   171
  | pattern_intersect (Split (n, xs', pat1)) (Split (m, ys, pat2)) =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   172
    Split (if n <= 0 then m else n, 
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   173
      intersect_consts xs' ys pat1 pat2,
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   174
      pattern_intersect pat1 pat2)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   175
and
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   176
    intersect_consts xs [] _ default2 = map (apsnd (fn pat => pattern_intersect pat default2)) xs
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   177
  | intersect_consts xs ((ctr, pat2) :: ys) default1 default2 = case find_ctr ctr xs of
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   178
    (SOME (ctr, pat1), xs') => 
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   179
      (ctr, pattern_merge (pattern_merge (pattern_intersect pat1 pat2) (pattern_intersect default1 pat2))
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   180
              (pattern_intersect pat1 default2)) ::
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   181
      intersect_consts xs' ys default1 default2
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   182
    | (NONE, xs') => (ctr, pattern_intersect default1 pat2) :: (intersect_consts xs' ys default1 default2)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   183
        
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   184
fun pattern_lookup _ Wildcard = Wildcard
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   185
  | pattern_lookup _ Value = Value
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   186
  | pattern_lookup [] (Split (n, xs, pat)) = 
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   187
    Split (n, map (apsnd (pattern_lookup [])) xs, pattern_lookup [] pat)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   188
  | pattern_lookup (term :: terms) (Split (n, xs, pat)) =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   189
  let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   190
    val (ctr, args) = strip_comb term
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   191
    fun map_ctr (term, pat) =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   192
      let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   193
        val args = term |> dest_Const |> snd |> binder_types |> map (fn T => Free ("x", T))
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   194
      in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   195
        pattern_lookup args pat
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   196
      end
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   197
  in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   198
    if is_Const ctr then
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   199
       case find_ctr ctr xs of (SOME (_, pat'), _) => 
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   200
           pattern_lookup terms (pattern_merge (pattern_lookup args pat') (pattern_lookup [] pat))
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   201
         | (NONE, _) => pattern_lookup terms pat
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   202
    else if length xs < n orelse n <= 0 then
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   203
      pattern_lookup terms pat
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   204
    else pattern_lookup terms
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   205
      (pattern_merge
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   206
        (fold pattern_intersect (map map_ctr (tl xs)) (map_ctr (hd xs)))
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   207
        (pattern_lookup [] pat))
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   208
  end;
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   209
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   210
fun pattern_contains terms pat = case pattern_lookup terms pat of
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   211
    Wildcard => false
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   212
  | Value => true
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   213
  | Split _ => raise Match;
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   214
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   215
fun pattern_create _ [] = Wildcard
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   216
  | pattern_create ctr_count (term :: terms) = 
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   217
    let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   218
      val (ctr, args) = strip_comb term
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   219
    in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   220
      if is_Const ctr then
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   221
        Split (ctr_count ctr, [(ctr, pattern_create ctr_count (args @ terms))], Wildcard)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   222
      else Split (0, [], pattern_create ctr_count terms)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   223
    end;
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   224
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   225
fun pattern_insert ctr_count terms pat =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   226
  let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   227
    fun new_pattern terms = pattern_insert ctr_count terms (pattern_create ctr_count terms)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   228
    fun aux _ false Wildcard = Wildcard
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   229
      | aux terms true Wildcard = if null terms then Value else new_pattern terms
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   230
      | aux _ _ Value = Value
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   231
      | aux terms modify (Split (n, xs', pat)) =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   232
      let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   233
        val unmodified = (n, map (apsnd (aux [] false)) xs', aux [] false pat)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   234
      in case terms of [] => Split unmodified
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   235
        | term :: terms =>
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   236
        let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   237
          val (ctr, args) = strip_comb term
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   238
          val (m, ys, pat') = unmodified
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   239
        in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   240
          if is_Const ctr
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   241
            then case find_ctr ctr xs' of
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   242
              (SOME (ctr, pat''), xs) =>
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   243
                Split (m, (ctr, aux (args @ terms) modify pat'') :: map (apsnd (aux [] false)) xs, pat')
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   244
              | (NONE, _) => if modify
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   245
                then if m <= 0
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   246
                  then Split (ctr_count ctr, (ctr, new_pattern (args @ terms)) :: ys, pat')
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   247
                  else Split (m, (ctr, new_pattern (args @ terms)) :: ys, pat')
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   248
                else Split unmodified
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   249
            else Split (m, ys, aux terms modify pat)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   250
        end
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   251
      end
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   252
  in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   253
    aux terms true pat
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   254
  end;
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   255
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   256
val pattern_empty = Wildcard;
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   257
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   258
fun replace_frees lhss rhss typ_list ctxt =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   259
  let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   260
    fun replace_frees_once (lhs, rhs) ctxt =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   261
      let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   262
        val add_frees_list = fold_rev Term.add_frees
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   263
        val frees = add_frees_list lhs []
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   264
        val (new_frees, ctxt1) = (Ctr_Sugar_Util.mk_Frees "x" (map snd frees) ctxt)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   265
        val (new_frees1, ctxt2) =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   266
          let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   267
            val (dest_frees, types) = split_list (map dest_Free new_frees)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   268
            val (new_frees, ctxt2) = Variable.variant_fixes dest_frees ctxt1
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   269
          in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   270
            (map Free (new_frees ~~ types), ctxt2)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   271
          end
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   272
        val dict = frees ~~ new_frees1
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   273
        fun free_map_fun (s, T) =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   274
          case AList.lookup (op =) dict (s, T) of
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   275
              NONE => Free (s, T)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   276
            | SOME x => x
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   277
        val map_fun = fold_term Const free_map_fun Var Bound Abs (op $)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   278
      in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   279
        ((map map_fun lhs, map_fun rhs), ctxt2)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   280
      end
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   281
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   282
    fun variant_fixes (def_frees, ctxt) =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   283
      let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   284
        val (dest_frees, types) = split_list (map dest_Free def_frees)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   285
        val (def_frees, ctxt1) = Variable.variant_fixes dest_frees ctxt
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   286
      in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   287
        (map Free (def_frees ~~ types), ctxt1)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   288
      end
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   289
    val (def_frees, ctxt1) = variant_fixes (Ctr_Sugar_Util.mk_Frees "x" typ_list ctxt)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   290
    val (rhs_frees, ctxt2) = variant_fixes (Ctr_Sugar_Util.mk_Frees "x" typ_list ctxt1)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   291
    val (case_args, ctxt3) = variant_fixes (Ctr_Sugar_Util.mk_Frees "x"
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   292
      (map fastype_of (hd lhss)) ctxt2)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   293
    val (new_terms1, ctxt4) = fold_map replace_frees_once (lhss ~~ rhss) ctxt3
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   294
    val (lhss1, rhss1) = split_list new_terms1
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   295
  in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   296
    (lhss1, rhss1, def_frees ~~ rhs_frees, case_args, ctxt4)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   297
  end;
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   298
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   299
fun add_names_in_type (Type (name, Ts)) = 
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   300
    List.foldr (op o) (Symtab.update (name, ())) (map add_names_in_type Ts)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   301
  | add_names_in_type (TFree _) = I
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   302
  | add_names_in_type (TVar _) = I
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   303
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   304
fun add_names_in_term (Const (_, T)) = add_names_in_type T
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   305
  | add_names_in_term (Free (_, T)) = add_names_in_type T
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   306
  | add_names_in_term (Var (_, T)) = add_names_in_type T
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   307
  | add_names_in_term (Bound _) = I
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   308
  | add_names_in_term (Abs (_, T, body)) =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   309
    add_names_in_type T o add_names_in_term body
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   310
  | add_names_in_term (t1 $ t2) = add_names_in_term t1 o add_names_in_term t2
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   311
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   312
fun add_type_names terms =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   313
  fold (fn term => fn f => add_names_in_term term o f) terms I
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   314
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   315
fun get_split_theorems ctxt =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   316
  Symtab.keys
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   317
  #> map_filter (Ctr_Sugar.ctr_sugar_of ctxt)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   318
  #> map #split;
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   319
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   320
fun match (Const (s1, _)) (Const (s2, _)) = if s1 = s2 then SOME I else NONE
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   321
  | match (Free y) x = SOME (fn z => if z = Free y then x else z)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   322
  | match (pat1 $ pattern2) (t1 $ t2) =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   323
    (case (match pat1 t1, match pattern2 t2) of
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   324
       (SOME f, SOME g) => SOME (f o g)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   325
       | _ => NONE
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   326
     )
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   327
  | match _ _ = NONE;
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   328
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   329
fun match_all patterns terms =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   330
  let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   331
    fun combine _ NONE = NONE
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   332
      | combine (f_opt, f_opt') (SOME g) = 
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   333
        case match f_opt f_opt' of SOME f => SOME (f o g) | _ => NONE
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   334
  in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   335
    fold_rev combine (patterns ~~ terms) (SOME I)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   336
  end
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   337
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   338
fun matches (Const (s1, _)) (Const (s2, _)) = s1 = s2
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   339
  | matches (Free _) _ = true 
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   340
  | matches (pat1 $ pat2) (t1 $ t2) = matches pat1 t1 andalso matches pat2 t2
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   341
  | matches _ _ = false;
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   342
fun matches_all patterns terms = forall (uncurry matches) (patterns ~~ terms)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   343
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   344
fun terms_to_case_at ctr_count ctxt (fun_t : term) (default_lhs : term list)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   345
    (pos, (lazy_case_arg, rhs_free))
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   346
    ((lhss : term list list), (rhss : term list), type_name_fun) =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   347
  let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   348
    fun abort t =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   349
      let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   350
        val fun_name = head_of t |> dest_Const |> fst
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   351
        val msg = "Missing pattern in " ^ fun_name ^ "."
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   352
      in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   353
        mk_abort msg t
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   354
      end;
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   355
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   356
    (* Step 1 : Eliminate lazy pattern *)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   357
    fun replace_pat_at (n, tcos) pat pats =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   358
      let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   359
        fun map_at _ _ [] = raise Empty
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   360
          | map_at n f (x :: xs) = if n > 0
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   361
            then apfst (cons x) (map_at (n - 1) f xs)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   362
            else apfst (fn x => x :: xs) (f x)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   363
        fun replace [] pat term = (pat, term)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   364
          | replace ((s1, n) :: tcos) pat term =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   365
            let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   366
              val (ctr, args) = strip_comb term
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   367
            in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   368
              case ctr of Const (s2, _) =>
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   369
                  if s1 = s2
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   370
                  then apfst (pair ctr #> list_comb) (map_at n (replace tcos pat) args)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   371
                  else (term, rhs_free)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   372
                | _ => (term, rhs_free)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   373
            end
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   374
        val (part1, (old_pat, part2)) = chop n pats ||> (fn xs => (hd xs, tl xs))
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   375
        val (new_pat, old_pat1) = replace tcos pat old_pat
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   376
      in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   377
        (part1 @ [new_pat] @ part2, old_pat1)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   378
      end                               
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   379
    val (lhss1, lazy_pats) = map (replace_pat_at pos lazy_case_arg) lhss
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   380
      |> split_list
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   381
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   382
    (* Step 2 : Split patterns *)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   383
    fun split equs =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   384
      let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   385
        fun merge_pattern (Const (s1, T1), Const (s2, _)) =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   386
            if s1 = s2 then SOME (Const (s1, T1)) else NONE
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   387
          | merge_pattern (t, Free _) = SOME t
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   388
          | merge_pattern (Free _, t) = SOME t
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   389
          | merge_pattern (t1l $ t1r, t2l $ t2r) =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   390
            (case (merge_pattern (t1l, t2l), merge_pattern (t1r, t2r)) of
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   391
              (SOME t1, SOME t2) => SOME (t1 $ t2)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   392
              | _ => NONE)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   393
          | merge_pattern _ = NONE
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   394
        fun merge_patterns pats1 pats2 = case (pats1, pats2) of
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   395
          ([], []) => SOME []
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   396
          | (x :: xs, y :: ys) =>
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   397
            (case (merge_pattern (x, y), merge_patterns xs ys) of
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   398
              (SOME x, SOME xs) => SOME (x :: xs)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   399
              | _ => NONE
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   400
            )
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   401
          | _ => raise Match
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   402
        fun merge_insert ((lhs1, case_pat), _) [] =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   403
            [(lhs1, pattern_empty |> pattern_insert ctr_count [case_pat])]
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   404
          | merge_insert ((lhs1, case_pat), rhs) ((lhs2, pat) :: pats) =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   405
            let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   406
              val pats = merge_insert ((lhs1, case_pat), rhs) pats
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   407
              val (first_equ_needed, new_lhs) = case merge_patterns lhs1 lhs2 of
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   408
                SOME new_lhs => (not (pattern_contains [case_pat] pat), new_lhs)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   409
                | NONE => (false, lhs2)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   410
              val second_equ_needed = not (matches_all lhs1 lhs2)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   411
                orelse not first_equ_needed
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   412
              val first_equ = if first_equ_needed
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   413
                then [(new_lhs, pattern_insert ctr_count [case_pat] pat)]
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   414
                else []
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   415
              val second_equ = if second_equ_needed
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   416
                then [(lhs2, pat)]
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   417
                else []
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   418
            in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   419
              first_equ @ second_equ @ pats
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   420
            end
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   421
        in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   422
          (fold merge_insert equs []
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   423
            |> split_list
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   424
            |> fst) @ [default_lhs]
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   425
        end
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   426
    val lhss2 = split ((lhss1 ~~ lazy_pats) ~~ rhss)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   427
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   428
    (* Step 3 : Remove redundant patterns *)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   429
    fun remove_redundant_lhs lhss =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   430
      let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   431
        fun f lhs pat = if pattern_contains lhs pat
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   432
          then ((lhs, false), pat)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   433
          else ((lhs, true), pattern_insert ctr_count lhs pat)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   434
      in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   435
        fold_map f lhss pattern_empty
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   436
        |> fst
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   437
        |> filter snd
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   438
        |> map fst
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   439
      end
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   440
    fun remove_redundant_rhs rhss =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   441
      let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   442
        fun f (lhs, rhs) pat = if pattern_contains [lhs] pat
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   443
          then (((lhs, rhs), false), pat)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   444
          else (((lhs, rhs), true), pattern_insert ctr_count [lhs] pat)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   445
      in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   446
        map fst (filter snd (fold_map f rhss pattern_empty |> fst))
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   447
      end
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   448
    val lhss3 = remove_redundant_lhs lhss2
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   449
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   450
    (* Step 4 : Compute right hand side *)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   451
    fun subs_fun f = fold_term
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   452
      Const
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   453
      (f o Free)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   454
      Var
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   455
      Bound
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   456
      Abs
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   457
      (fn (x, y) => f x $ f y)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   458
    fun find_rhss lhs =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   459
      let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   460
        fun f (lhs1, (pat, rhs)) = 
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   461
          case match_all lhs1 lhs of NONE => NONE
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   462
          | SOME f => SOME (pat, subs_fun f rhs)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   463
      in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   464
        remove_redundant_rhs
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   465
          (map_filter f (lhss1 ~~ (lazy_pats ~~ rhss)) @
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   466
            [(lazy_case_arg, list_comb (fun_t, lhs) |> abort)]
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   467
          )
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   468
      end
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   469
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   470
    (* Step 5 : make_case of right hand side *)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   471
    fun make_case ctxt case_arg cases = case cases of
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   472
      [(Free x, rhs)] => subs_fun (fn y => if y = Free x then case_arg else y) rhs
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   473
      | _ => Case_Translation.make_case
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   474
        ctxt
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   475
        Case_Translation.Warning
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   476
        Name.context
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   477
        case_arg
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   478
        cases
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   479
    val type_name_fun = add_type_names lazy_pats o type_name_fun
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   480
    val rhss3 = map ((make_case ctxt lazy_case_arg) o find_rhss) lhss3
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   481
  in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   482
    (lhss3, rhss3, type_name_fun)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   483
  end;
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   484
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   485
fun terms_to_case ctxt ctr_count (head : term) (lhss : term list list)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   486
    (rhss : term list) (typ_list : typ list) (poss : (int * (string * int) list) list) =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   487
  let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   488
    val (lhss1, rhss1, def_frees, case_args, ctxt1) = replace_frees lhss rhss typ_list ctxt
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   489
    val exec_list = poss ~~ def_frees
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   490
    val (lhss2, rhss2, type_name_fun) = fold_rev
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   491
      (terms_to_case_at ctr_count ctxt1 head case_args) exec_list (lhss1, rhss1, I)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   492
    fun make_eq_term (lhss, rhs) = (list_comb (head, lhss), rhs)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   493
      |> HOLogic.mk_eq
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   494
      |> HOLogic.mk_Trueprop
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   495
  in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   496
    (map make_eq_term (lhss2 ~~ rhss2),
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   497
      get_split_theorems ctxt1 (type_name_fun Symtab.empty),
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   498
      ctxt1)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   499
  end;
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   500
69568
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   501
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   502
fun build_case_t elimination_strategy ctr_count head lhss rhss ctxt =
68155
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   503
  let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   504
    val num_eqs = length lhss
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   505
    val _ = if length rhss = num_eqs andalso num_eqs > 0 then ()
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   506
      else raise Fail
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   507
        ("expected same number of left-hand sides as right-hand sides\n"
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   508
          ^ "and at least one equation")
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   509
    val n = length (hd lhss)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   510
    val _ = if forall (fn m => length m = n) lhss then ()
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   511
      else raise Fail "expected equal number of arguments"
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   512
69568
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   513
    fun to_coordinates (n, ts) = 
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   514
      case elimination_strategy ctxt ts of
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   515
          [] => NONE
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   516
        | (tco :: tcos) => SOME (n, fold term_coordinate_merge tcos tco |> coordinates_to_list)
68155
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   517
    fun add_T (n, xss) = map (fn (T, xs) => (T, (n, xs))) xss
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   518
    val (typ_list, poss) = lhss
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   519
      |> Ctr_Sugar_Util.transpose
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   520
      |> map_index to_coordinates
68301
fb5653a7a879 prefer existing operation;
wenzelm
parents: 68155
diff changeset
   521
      |> map_filter (Option.map add_T)
68155
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   522
      |> flat
69568
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 68301
diff changeset
   523
      |> split_list
68155
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   524
  in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   525
    if null poss then ([], [], ctxt)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   526
    else terms_to_case ctxt (dest_Const #> ctr_count) head lhss rhss typ_list poss
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   527
  end;
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   528
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   529
fun tac ctxt {splits, intros, defs} =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   530
  let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   531
    val split_and_subst = 
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   532
      split_tac ctxt splits 
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   533
      THEN' REPEAT_ALL_NEW (
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   534
        resolve_tac ctxt [@{thm conjI}, @{thm allI}]
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   535
        ORELSE'
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   536
        (resolve_tac ctxt [@{thm impI}] THEN' hyp_subst_tac_thin true ctxt))
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   537
  in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   538
    (REPEAT_ALL_NEW split_and_subst ORELSE' K all_tac)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   539
    THEN' (K (Local_Defs.unfold_tac ctxt [@{thm missing_pattern_match_def}]))
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   540
    THEN' (K (Local_Defs.unfold_tac ctxt defs))
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   541
    THEN_ALL_NEW (SOLVED' (resolve_tac ctxt (@{thm refl} :: intros)))
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   542
  end;
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   543
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   544
fun to_case _ _ _ [] = NONE
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   545
  | to_case ctxt replace_ctr ctr_count ths =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   546
    let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   547
      val strip_eq = Thm.prop_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   548
      fun import [] ctxt = ([], ctxt)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   549
        | import (thm :: thms) ctxt =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   550
          let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   551
            val fun_ct = strip_eq #> fst #> head_of #> Logic.mk_term #> Thm.cterm_of ctxt
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   552
            val ct = fun_ct thm
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   553
            val cts = map fun_ct thms
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   554
            val pairs = map (fn s => (s,ct)) cts
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   555
            val thms' = map (fn (th,p) => Thm.instantiate (Thm.match p) th) (thms ~~ pairs)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   556
          in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   557
            Variable.import true (thm :: thms') ctxt |> apfst snd
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   558
          end
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   559
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   560
      val (iths, ctxt') = import ths ctxt
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   561
      val head = hd iths |> strip_eq |> fst |> head_of
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   562
      val eqs = map (strip_eq #> apfst (snd o strip_comb)) iths
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   563
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   564
      fun hide_rhs ((pat, rhs), name) lthy =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   565
        let
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   566
          val frees = fold Term.add_frees pat []
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   567
          val abs_rhs = fold absfree frees rhs
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   568
          val (f, def, lthy') = case lthy
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   569
            |> Local_Defs.define [((Binding.name name, NoSyn), (Binding.empty_atts, abs_rhs))] of
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   570
              ([(f, (_, def))], lthy') => (f, def, lthy')
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   571
              | _ => raise Match
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   572
        in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   573
          ((list_comb (f, map Free (rev frees)), def), lthy')
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   574
        end
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   575
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   576
      val rhs_names = Name.invent (Variable.names_of ctxt') "rhs" (length eqs)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   577
      val ((def_ts, def_thms), ctxt2) =
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   578
        fold_map hide_rhs (eqs ~~ rhs_names) ctxt' |> apfst split_list
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   579
      val (ts, split_thms, ctxt3) = build_case_t replace_ctr ctr_count head
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   580
        (map fst eqs) def_ts ctxt2
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   581
      fun mk_thm t = Goal.prove ctxt3 [] [] t
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   582
          (fn {context=ctxt, ...} => tac ctxt {splits=split_thms, intros=ths, defs=def_thms} 1)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   583
    in
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   584
      if null ts then NONE
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   585
      else
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   586
        ts
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   587
        |> map mk_thm
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   588
        |> Proof_Context.export ctxt3 ctxt
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   589
        |> map (Goal.norm_result ctxt)
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   590
        |> SOME
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   591
    end;
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   592
8b50f29a1992 new tool Code_Lazy
Andreas Lochbihler
parents:
diff changeset
   593
end