src/HOL/HOLCF/Tools/fixrec.ML
author wenzelm
Fri, 16 Dec 2011 10:52:35 +0100
changeset 45897 65cef0298158
parent 45787 9fcaf2557c59
child 46895 de5cfda8b2de
permissions -rw-r--r--
clarified modules that contribute to datatype package;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42151
4da4fc77664b tuned headers;
wenzelm
parents: 41296
diff changeset
     1
(*  Title:      HOL/HOLCF/Tools/fixrec.ML
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     2
    Author:     Amber Telfer and Brian Huffman
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     3
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     4
Recursive function definition package for HOLCF.
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     5
*)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     6
31738
7b9b9ba532ca discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31177
diff changeset
     7
signature FIXREC =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     8
sig
40041
1f09b4c7b85e replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents: 40036
diff changeset
     9
  val add_fixrec: (binding * typ option * mixfix) list
1f09b4c7b85e replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents: 40036
diff changeset
    10
    -> (bool * (Attrib.binding * term)) list -> local_theory -> local_theory
1f09b4c7b85e replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents: 40036
diff changeset
    11
  val add_fixrec_cmd: (binding * string option * mixfix) list
1f09b4c7b85e replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents: 40036
diff changeset
    12
    -> (bool * (Attrib.binding * string)) list -> local_theory -> local_theory
30131
6be1be402ef0 use TheoryData to keep track of pattern match combinators
huffman
parents: 29585
diff changeset
    13
  val add_matchers: (string * string) list -> theory -> theory
33425
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
    14
  val fixrec_simp_tac: Proof.context -> int -> tactic
30131
6be1be402ef0 use TheoryData to keep track of pattern match combinators
huffman
parents: 29585
diff changeset
    15
  val setup: theory -> theory
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    16
end
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    17
41296
6aaf80ea9715 switch to transparent ascription, to avoid warning messages
huffman
parents: 40832
diff changeset
    18
structure Fixrec : FIXREC =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    19
struct
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    20
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    21
open HOLCF_Library
35527
f4282471461d fixrec and repdef modules import holcf_library
huffman
parents: 35525
diff changeset
    22
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    23
infixr 6 ->>
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    24
infix -->>
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    25
infix 9 `
35527
f4282471461d fixrec and repdef modules import holcf_library
huffman
parents: 35525
diff changeset
    26
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    27
val def_cont_fix_eq = @{thm def_cont_fix_eq}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    28
val def_cont_fix_ind = @{thm def_cont_fix_ind}
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    29
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    30
fun fixrec_err s = error ("fixrec definition error:\n" ^ s)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    31
30132
243a05a67c41 avoid using legacy type inference
huffman
parents: 30131
diff changeset
    32
(*************************************************************************)
243a05a67c41 avoid using legacy type inference
huffman
parents: 30131
diff changeset
    33
(***************************** building types ****************************)
243a05a67c41 avoid using legacy type inference
huffman
parents: 30131
diff changeset
    34
(*************************************************************************)
243a05a67c41 avoid using legacy type inference
huffman
parents: 30131
diff changeset
    35
33401
fc43fa403a69 add fixrec support for HOL pair constructor patterns
huffman
parents: 33004
diff changeset
    36
local
fc43fa403a69 add fixrec support for HOL pair constructor patterns
huffman
parents: 33004
diff changeset
    37
35525
fa231b86cb1e proper names for types cfun, sprod, ssum
huffman
parents: 33766
diff changeset
    38
fun binder_cfun (Type(@{type_name cfun},[T, U])) = T :: binder_cfun U
33401
fc43fa403a69 add fixrec support for HOL pair constructor patterns
huffman
parents: 33004
diff changeset
    39
  | binder_cfun (Type(@{type_name "fun"},[T, U])) = T :: binder_cfun U
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    40
  | binder_cfun _   =  []
33401
fc43fa403a69 add fixrec support for HOL pair constructor patterns
huffman
parents: 33004
diff changeset
    41
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
    42
fun body_cfun (Type(@{type_name cfun},[_, U])) = body_cfun U
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
    43
  | body_cfun (Type(@{type_name "fun"},[_, U])) = body_cfun U
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    44
  | body_cfun T   =  T
33401
fc43fa403a69 add fixrec support for HOL pair constructor patterns
huffman
parents: 33004
diff changeset
    45
fc43fa403a69 add fixrec support for HOL pair constructor patterns
huffman
parents: 33004
diff changeset
    46
in
fc43fa403a69 add fixrec support for HOL pair constructor patterns
huffman
parents: 33004
diff changeset
    47
35527
f4282471461d fixrec and repdef modules import holcf_library
huffman
parents: 35525
diff changeset
    48
fun matcherT (T, U) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    49
  body_cfun T ->> (binder_cfun T -->> U) ->> U
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30485
diff changeset
    50
33401
fc43fa403a69 add fixrec support for HOL pair constructor patterns
huffman
parents: 33004
diff changeset
    51
end
30132
243a05a67c41 avoid using legacy type inference
huffman
parents: 30131
diff changeset
    52
243a05a67c41 avoid using legacy type inference
huffman
parents: 30131
diff changeset
    53
(*************************************************************************)
243a05a67c41 avoid using legacy type inference
huffman
parents: 30131
diff changeset
    54
(***************************** building terms ****************************)
243a05a67c41 avoid using legacy type inference
huffman
parents: 30131
diff changeset
    55
(*************************************************************************)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    56
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    57
val mk_trp = HOLogic.mk_Trueprop
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    58
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    59
(* splits a cterm into the right and lefthand sides of equality *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    60
fun dest_eqs t = HOLogic.dest_eq (HOLogic.dest_Trueprop t)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    61
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    62
(* similar to Thm.head_of, but for continuous application *)
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
    63
fun chead_of (Const(@{const_name Rep_cfun},_)$f$_) = chead_of f
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    64
  | chead_of u = u
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    65
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    66
infix 1 === val (op ===) = HOLogic.mk_eq
30132
243a05a67c41 avoid using legacy type inference
huffman
parents: 30131
diff changeset
    67
243a05a67c41 avoid using legacy type inference
huffman
parents: 30131
diff changeset
    68
fun mk_mplus (t, u) =
243a05a67c41 avoid using legacy type inference
huffman
parents: 30131
diff changeset
    69
  let val mT = Term.fastype_of t
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    70
  in Const(@{const_name Fixrec.mplus}, mT ->> mT ->> mT) ` t ` u end
30132
243a05a67c41 avoid using legacy type inference
huffman
parents: 30131
diff changeset
    71
243a05a67c41 avoid using legacy type inference
huffman
parents: 30131
diff changeset
    72
fun mk_run t =
37097
476016cbf8b3 for functions with only variable patterns, fixrec definitions no longer use Fixrec.return/Fixrec.run
huffman
parents: 37096
diff changeset
    73
  let
476016cbf8b3 for functions with only variable patterns, fixrec definitions no longer use Fixrec.return/Fixrec.run
huffman
parents: 37096
diff changeset
    74
    val mT = Term.fastype_of t
476016cbf8b3 for functions with only variable patterns, fixrec definitions no longer use Fixrec.return/Fixrec.run
huffman
parents: 37096
diff changeset
    75
    val T = dest_matchT mT
476016cbf8b3 for functions with only variable patterns, fixrec definitions no longer use Fixrec.return/Fixrec.run
huffman
parents: 37096
diff changeset
    76
    val run = Const(@{const_name Fixrec.run}, mT ->> T)
476016cbf8b3 for functions with only variable patterns, fixrec definitions no longer use Fixrec.return/Fixrec.run
huffman
parents: 37096
diff changeset
    77
  in
476016cbf8b3 for functions with only variable patterns, fixrec definitions no longer use Fixrec.return/Fixrec.run
huffman
parents: 37096
diff changeset
    78
    case t of
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40096
diff changeset
    79
      Const(@{const_name Rep_cfun}, _) $
37108
00f13d3ad474 rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents: 37097
diff changeset
    80
        Const(@{const_name Fixrec.succeed}, _) $ u => u
37097
476016cbf8b3 for functions with only variable patterns, fixrec definitions no longer use Fixrec.return/Fixrec.run
huffman
parents: 37096
diff changeset
    81
    | _ => run ` t
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    82
  end
30132
243a05a67c41 avoid using legacy type inference
huffman
parents: 30131
diff changeset
    83
31095
b79d140f6d0b simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents: 31023
diff changeset
    84
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    85
(*************************************************************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    86
(************* fixed-point definitions and unfolding theorems ************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    87
(*************************************************************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    88
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 33507
diff changeset
    89
structure FixrecUnfoldData = Generic_Data
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 33507
diff changeset
    90
(
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    91
  type T = thm Symtab.table
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    92
  val empty = Symtab.empty
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    93
  val extend = I
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    94
  fun merge data : T = Symtab.merge (K true) data
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    95
)
33425
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
    96
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
    97
local
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
    98
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
    99
fun name_of (Const (n, _)) = n
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   100
  | name_of (Free (n, _)) = n
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   101
  | name_of t = raise TERM ("Fixrec.add_unfold: lhs not a constant", [t])
33425
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   102
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   103
val lhs_name =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   104
  name_of o head_of o fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of
33425
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   105
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   106
in
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   107
33442
5d78b2bd27de made SML/NJ happy;
wenzelm
parents: 33430
diff changeset
   108
val add_unfold : attribute =
33425
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   109
  Thm.declaration_attribute
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   110
    (fn th => FixrecUnfoldData.map (Symtab.insert (K true) (lhs_name th, th)))
33425
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   111
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   112
end
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   113
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 30157
diff changeset
   114
fun add_fixdefs
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 30157
diff changeset
   115
  (fixes : ((binding * typ) * mixfix) list)
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 30157
diff changeset
   116
  (spec : (Attrib.binding * term) list)
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 30157
diff changeset
   117
  (lthy : local_theory) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   118
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42151
diff changeset
   119
    val thy = Proof_Context.theory_of lthy
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   120
    val names = map (Binding.name_of o fst o fst) fixes
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   121
    val all_names = space_implode "_" names
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   122
    val (lhss, rhss) = ListPair.unzip (map (dest_eqs o snd) spec)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   123
    val functional = lambda_tuple lhss (mk_tuple rhss)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   124
    val fixpoint = mk_fix (mk_cabs functional)
36996
63fadc0a33db more informative error message for fixrec when continuity proof fails
huffman
parents: 36960
diff changeset
   125
31095
b79d140f6d0b simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents: 31023
diff changeset
   126
    val cont_thm =
36996
63fadc0a33db more informative error message for fixrec when continuity proof fails
huffman
parents: 36960
diff changeset
   127
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   128
        val prop = mk_trp (mk_cont functional)
37096
7b74b4a734fd simplify fixrec continuity tactic
huffman
parents: 37084
diff changeset
   129
        fun err _ = error (
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   130
          "Continuity proof failed please check that cont2cont rules\n" ^
37097
476016cbf8b3 for functions with only variable patterns, fixrec definitions no longer use Fixrec.return/Fixrec.run
huffman
parents: 37096
diff changeset
   131
          "or simp rules are configured for all non-HOLCF constants.\n" ^
36996
63fadc0a33db more informative error message for fixrec when continuity proof fails
huffman
parents: 36960
diff changeset
   132
          "The error occurred for the goal statement:\n" ^
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   133
          Syntax.string_of_term lthy prop)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   134
        val rules = Cont2ContData.get lthy
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   135
        val fast_tac = SOLVED' (REPEAT_ALL_NEW (match_tac rules))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   136
        val slow_tac = SOLVED' (simp_tac (simpset_of lthy))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   137
        val tac = fast_tac 1 ORELSE slow_tac 1 ORELSE err
36996
63fadc0a33db more informative error message for fixrec when continuity proof fails
huffman
parents: 36960
diff changeset
   138
      in
63fadc0a33db more informative error message for fixrec when continuity proof fails
huffman
parents: 36960
diff changeset
   139
        Goal.prove lthy [] [] prop (K tac)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   140
      end
31095
b79d140f6d0b simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents: 31023
diff changeset
   141
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   142
    fun one_def (Free(n,_)) r =
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30280
diff changeset
   143
          let val b = Long_Name.base_name n
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 30157
diff changeset
   144
          in ((Binding.name (b^"_def"), []), r) end
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   145
      | one_def _ _ = fixrec_err "fixdefs: lhs not of correct form"
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   146
    fun defs [] _ = []
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   147
      | defs (l::[]) r = [one_def l r]
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   148
      | defs (l::ls) r = one_def l (mk_fst r) :: defs ls (mk_snd r)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   149
    val fixdefs = defs lhss fixpoint
35772
ea0ac5538c53 avoid unnecessary primed variable names
huffman
parents: 35770
diff changeset
   150
    val (fixdef_thms : (term * (string * thm)) list, lthy) = lthy
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   151
      |> fold_map Local_Theory.define (map (apfst fst) fixes ~~ fixdefs)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   152
    fun pair_equalI (thm1, thm2) = @{thm Pair_equalI} OF [thm1, thm2]
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   153
    val tuple_fixdef_thm = foldr1 pair_equalI (map (snd o snd) fixdef_thms)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   154
    val P = Var (("P", 0), map Term.fastype_of lhss ---> HOLogic.boolT)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   155
    val predicate = lambda_tuple lhss (list_comb (P, lhss))
31095
b79d140f6d0b simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents: 31023
diff changeset
   156
    val tuple_induct_thm = (def_cont_fix_ind OF [tuple_fixdef_thm, cont_thm])
b79d140f6d0b simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents: 31023
diff changeset
   157
      |> Drule.instantiate' [] [SOME (Thm.cterm_of thy predicate)]
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   158
      |> Local_Defs.unfold lthy @{thms split_paired_all split_conv split_strict}
31095
b79d140f6d0b simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents: 31023
diff changeset
   159
    val tuple_unfold_thm = (def_cont_fix_eq OF [tuple_fixdef_thm, cont_thm])
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   160
      |> Local_Defs.unfold lthy @{thms split_conv}
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   161
    fun unfolds [] _ = []
33425
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   162
      | unfolds (n::[]) thm = [(n, thm)]
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   163
      | unfolds (n::ns) thm = let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   164
          val thmL = thm RS @{thm Pair_eqD1}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   165
          val thmR = thm RS @{thm Pair_eqD2}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   166
        in (n, thmL) :: unfolds ns thmR end
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   167
    val unfold_thms = unfolds names tuple_unfold_thm
33425
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   168
    val induct_note : Attrib.binding * Thm.thm list =
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   169
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   170
        val thm_name = Binding.qualify true all_names (Binding.name "induct")
33425
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   171
      in
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   172
        ((thm_name, []), [tuple_induct_thm])
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   173
      end
33425
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   174
    fun unfold_note (name, thm) : Attrib.binding * Thm.thm list =
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   175
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   176
        val thm_name = Binding.qualify true name (Binding.name "unfold")
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   177
        val src = Attrib.internal (K add_unfold)
33425
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   178
      in
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   179
        ((thm_name, [src]), [thm])
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   180
      end
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   181
    val (_, lthy) = lthy
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   182
      |> fold_map Local_Theory.note (induct_note :: map unfold_note unfold_thms)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   183
  in
35772
ea0ac5538c53 avoid unnecessary primed variable names
huffman
parents: 35770
diff changeset
   184
    (lthy, names, fixdef_thms, map snd unfold_thms)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   185
  end
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   186
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   187
(*************************************************************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   188
(*********** monadic notation and pattern matching compilation ***********)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   189
(*************************************************************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   190
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33519
diff changeset
   191
structure FixrecMatchData = Theory_Data
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33519
diff changeset
   192
(
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   193
  type T = string Symtab.table
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   194
  val empty = Symtab.empty
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   195
  val extend = I
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   196
  fun merge data = Symtab.merge (K true) data
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   197
)
30131
6be1be402ef0 use TheoryData to keep track of pattern match combinators
huffman
parents: 29585
diff changeset
   198
6be1be402ef0 use TheoryData to keep track of pattern match combinators
huffman
parents: 29585
diff changeset
   199
(* associate match functions with pattern constants *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   200
fun add_matchers ms = FixrecMatchData.map (fold Symtab.update ms)
30131
6be1be402ef0 use TheoryData to keep track of pattern match combinators
huffman
parents: 29585
diff changeset
   201
30157
40919ebde2c9 add function taken_names
huffman
parents: 30132
diff changeset
   202
fun taken_names (t : term) : bstring list =
40919ebde2c9 add function taken_names
huffman
parents: 30132
diff changeset
   203
  let
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30280
diff changeset
   204
    fun taken (Const(a,_), bs) = insert (op =) (Long_Name.base_name a) bs
30157
40919ebde2c9 add function taken_names
huffman
parents: 30132
diff changeset
   205
      | taken (Free(a,_) , bs) = insert (op =) a bs
40919ebde2c9 add function taken_names
huffman
parents: 30132
diff changeset
   206
      | taken (f $ u     , bs) = taken (f, taken (u, bs))
40919ebde2c9 add function taken_names
huffman
parents: 30132
diff changeset
   207
      | taken (Abs(a,_,t), bs) = taken (t, insert (op =) a bs)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   208
      | taken (_         , bs) = bs
30157
40919ebde2c9 add function taken_names
huffman
parents: 30132
diff changeset
   209
  in
40919ebde2c9 add function taken_names
huffman
parents: 30132
diff changeset
   210
    taken (t, [])
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   211
  end
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   212
39805
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   213
(* builds a monadic term for matching a pattern *)
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   214
(* returns (rhs, free variable, used varnames) *)
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   215
fun compile_pat match_name pat rhs taken =
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   216
  let
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   217
    fun comp_pat p rhs taken =
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   218
      if is_Free p then (rhs, p, taken)
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   219
      else comp_con (fastype_of p) p rhs [] taken
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   220
    (* compiles a monadic term for a constructor pattern *)
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   221
    and comp_con T p rhs vs taken =
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   222
      case p of
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40096
diff changeset
   223
        Const(@{const_name Rep_cfun},_) $ f $ x =>
39805
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   224
          let val (rhs', v, taken') = comp_pat x rhs taken
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   225
          in comp_con T f rhs' (v::vs) taken' end
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   226
      | f $ x =>
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   227
          let val (rhs', v, taken') = comp_pat x rhs taken
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   228
          in comp_con T f rhs' (v::vs) taken' end
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   229
      | Const (c, cT) =>
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   230
          let
43324
2b47822868e4 discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents: 42361
diff changeset
   231
            val n = singleton (Name.variant_list taken) "v"
39805
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   232
            val v = Free(n, T)
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   233
            val m = Const(match_name c, matcherT (cT, fastype_of rhs))
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   234
            val k = big_lambdas vs rhs
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   235
          in
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   236
            (m`v`k, v, n::taken)
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   237
          end
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   238
      | _ => raise TERM ("fixrec: invalid pattern ", [p])
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   239
  in
16c53975ae1a clean up definition of compile_pat function
huffman
parents: 39804
diff changeset
   240
    comp_pat pat rhs taken
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   241
  end
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   242
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   243
(* builds a monadic term for matching a function definition pattern *)
40036
a81758e0394d simplify fixrec pattern match function
huffman
parents: 39806
diff changeset
   244
(* returns (constant, (vars, matcher)) *)
39804
b1cec1fcd95f rename some fixrec pattern-match compilation functions
huffman
parents: 39557
diff changeset
   245
fun compile_lhs match_name pat rhs vs taken =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   246
  case pat of
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40096
diff changeset
   247
    Const(@{const_name Rep_cfun}, _) $ f $ x =>
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   248
      let val (rhs', v, taken') = compile_pat match_name x rhs taken
39804
b1cec1fcd95f rename some fixrec pattern-match compilation functions
huffman
parents: 39557
diff changeset
   249
      in compile_lhs match_name f rhs' (v::vs) taken' end
40036
a81758e0394d simplify fixrec pattern match function
huffman
parents: 39806
diff changeset
   250
  | Free(_,_) => (pat, (vs, rhs))
a81758e0394d simplify fixrec pattern match function
huffman
parents: 39806
diff changeset
   251
  | Const(_,_) => (pat, (vs, rhs))
a81758e0394d simplify fixrec pattern match function
huffman
parents: 39806
diff changeset
   252
  | _ => fixrec_err ("invalid function pattern: "
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   253
                    ^ ML_Syntax.print_term pat)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   254
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 30157
diff changeset
   255
fun strip_alls t =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   256
  if Logic.is_all t then strip_alls (snd (Logic.dest_all t)) else t
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 30157
diff changeset
   257
39804
b1cec1fcd95f rename some fixrec pattern-match compilation functions
huffman
parents: 39557
diff changeset
   258
fun compile_eq match_name eq =
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 30157
diff changeset
   259
  let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   260
    val (lhs,rhs) = dest_eqs (Logic.strip_imp_concl (strip_alls eq))
30131
6be1be402ef0 use TheoryData to keep track of pattern match combinators
huffman
parents: 29585
diff changeset
   261
  in
39804
b1cec1fcd95f rename some fixrec pattern-match compilation functions
huffman
parents: 39557
diff changeset
   262
    compile_lhs match_name lhs (mk_succeed rhs) [] (taken_names eq)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   263
  end
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   264
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   265
(* this is the pattern-matching compiler function *)
39804
b1cec1fcd95f rename some fixrec pattern-match compilation functions
huffman
parents: 39557
diff changeset
   266
fun compile_eqs match_name eqs =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   267
  let
40036
a81758e0394d simplify fixrec pattern match function
huffman
parents: 39806
diff changeset
   268
    val (consts, matchers) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   269
      ListPair.unzip (map (compile_eq match_name) eqs)
40036
a81758e0394d simplify fixrec pattern match function
huffman
parents: 39806
diff changeset
   270
    val const =
a81758e0394d simplify fixrec pattern match function
huffman
parents: 39806
diff changeset
   271
        case distinct (op =) consts of
35903
0b43ff2d2e91 fix ML warnings in fixrec.ML
huffman
parents: 35780
diff changeset
   272
          [n] => n
45787
9fcaf2557c59 more error checking for fixrec
huffman
parents: 45592
diff changeset
   273
        | [] => fixrec_err "no defining equations for function"
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   274
        | _ => fixrec_err "all equations in block must define the same function"
40036
a81758e0394d simplify fixrec pattern match function
huffman
parents: 39806
diff changeset
   275
    val vars =
a81758e0394d simplify fixrec pattern match function
huffman
parents: 39806
diff changeset
   276
        case distinct (op = o pairself length) (map fst matchers) of
a81758e0394d simplify fixrec pattern match function
huffman
parents: 39806
diff changeset
   277
          [vars] => vars
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   278
        | _ => fixrec_err "all equations in block must have the same arity"
40036
a81758e0394d simplify fixrec pattern match function
huffman
parents: 39806
diff changeset
   279
    (* rename so all matchers use same free variables *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   280
    fun rename (vs, t) = Term.subst_free (filter_out (op =) (vs ~~ vars)) t
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   281
    val rhs = big_lambdas vars (mk_run (foldr1 mk_mplus (map rename matchers)))
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   282
  in
40036
a81758e0394d simplify fixrec pattern match function
huffman
parents: 39806
diff changeset
   283
    mk_trp (const === rhs)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   284
  end
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   285
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   286
(*************************************************************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   287
(********************** Proving associated theorems **********************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   288
(*************************************************************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   289
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   290
fun eta_tac i = CONVERSION Thm.eta_conversion i
33425
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   291
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   292
fun fixrec_simp_tac ctxt =
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   293
  let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   294
    val tab = FixrecUnfoldData.get (Context.Proof ctxt)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   295
    val ss = Simplifier.simpset_of ctxt
33425
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   296
    fun concl t =
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   297
      if Logic.is_all t then concl (snd (Logic.dest_all t))
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   298
      else HOLogic.dest_Trueprop (Logic.strip_imp_concl t)
33425
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   299
    fun tac (t, i) =
33430
c7dfeb7b0b0e better error handling for fixrec_simp
huffman
parents: 33427
diff changeset
   300
      let
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   301
        val (c, _) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   302
            (dest_Const o head_of o chead_of o fst o HOLogic.dest_eq o concl) t
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   303
        val unfold_thm = the (Symtab.lookup tab c)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   304
        val rule = unfold_thm RS @{thm ssubst_lhs}
33430
c7dfeb7b0b0e better error handling for fixrec_simp
huffman
parents: 33427
diff changeset
   305
      in
37081
3e5146b93218 simplify definition of eta_tac
huffman
parents: 37080
diff changeset
   306
        CHANGED (rtac rule i THEN eta_tac i THEN asm_simp_tac ss i)
33430
c7dfeb7b0b0e better error handling for fixrec_simp
huffman
parents: 33427
diff changeset
   307
      end
33425
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   308
  in
33505
03221db9df29 use try instead of handle
huffman
parents: 33430
diff changeset
   309
    SUBGOAL (fn ti => the_default no_tac (try tac ti))
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   310
  end
33425
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   311
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   312
(* proves a block of pattern matching equations as theorems, using unfold *)
32149
ef59550a55d3 renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
wenzelm
parents: 31740
diff changeset
   313
fun make_simps ctxt (unfold_thm, eqns : (Attrib.binding * term) list) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   314
  let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   315
    val ss = Simplifier.simpset_of ctxt
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   316
    val rule = unfold_thm RS @{thm ssubst_lhs}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   317
    val tac = rtac rule 1 THEN eta_tac 1 THEN asm_simp_tac ss 1
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   318
    fun prove_term t = Goal.prove ctxt [] [] t (K tac)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   319
    fun prove_eqn (bind, eqn_t) = (bind, prove_term eqn_t)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   320
  in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   321
    map prove_eqn eqns
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   322
  end
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   323
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   324
(*************************************************************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   325
(************************* Main fixrec function **************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   326
(*************************************************************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   327
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 30157
diff changeset
   328
local
45897
65cef0298158 clarified modules that contribute to datatype package;
wenzelm
parents: 45787
diff changeset
   329
(* code adapted from HOL/Tools/Datatype/primrec.ML *)
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 30157
diff changeset
   330
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 30157
diff changeset
   331
fun gen_fixrec
30485
99def5248e7f removed legacy_infer_term, legacy_infer_prop;
wenzelm
parents: 30364
diff changeset
   332
  prep_spec
40041
1f09b4c7b85e replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents: 40036
diff changeset
   333
  (raw_fixes : (binding * 'a option * mixfix) list)
1f09b4c7b85e replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents: 40036
diff changeset
   334
  (raw_spec' : (bool * (Attrib.binding * 'b)) list)
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 30157
diff changeset
   335
  (lthy : local_theory) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   336
  let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   337
    val (skips, raw_spec) = ListPair.unzip raw_spec'
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 30157
diff changeset
   338
    val (fixes : ((binding * typ) * mixfix) list,
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 30157
diff changeset
   339
         spec : (Attrib.binding * term) list) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   340
          fst (prep_spec raw_fixes raw_spec lthy)
45787
9fcaf2557c59 more error checking for fixrec
huffman
parents: 45592
diff changeset
   341
    val names = map (Binding.name_of o fst o fst) fixes
9fcaf2557c59 more error checking for fixrec
huffman
parents: 45592
diff changeset
   342
    fun check_head name =
9fcaf2557c59 more error checking for fixrec
huffman
parents: 45592
diff changeset
   343
        member (op =) names name orelse
9fcaf2557c59 more error checking for fixrec
huffman
parents: 45592
diff changeset
   344
        fixrec_err ("Illegal equation head. Expected " ^ commas_quote names)
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 30157
diff changeset
   345
    val chead_of_spec =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   346
      chead_of o fst o dest_eqs o Logic.strip_imp_concl o strip_alls o snd
45787
9fcaf2557c59 more error checking for fixrec
huffman
parents: 45592
diff changeset
   347
    fun name_of (Free (n, _)) = tap check_head n
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   348
      | name_of _ = fixrec_err ("unknown term")
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   349
    val all_names = map (name_of o chead_of_spec) spec
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 30157
diff changeset
   350
    fun block_of_name n =
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 30157
diff changeset
   351
      map_filter
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 30157
diff changeset
   352
        (fn (m,eq) => if m = n then SOME eq else NONE)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   353
        (all_names ~~ (spec ~~ skips))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   354
    val blocks = map block_of_name names
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 30157
diff changeset
   355
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42151
diff changeset
   356
    val matcher_tab = FixrecMatchData.get (Proof_Context.theory_of lthy)
30131
6be1be402ef0 use TheoryData to keep track of pattern match combinators
huffman
parents: 29585
diff changeset
   357
    fun match_name c =
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 30157
diff changeset
   358
      case Symtab.lookup matcher_tab c of SOME m => m
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   359
        | NONE => fixrec_err ("unknown pattern constructor: " ^ c)
30131
6be1be402ef0 use TheoryData to keep track of pattern match combinators
huffman
parents: 29585
diff changeset
   360
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   361
    val matches = map (compile_eqs match_name) (map (map (snd o fst)) blocks)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   362
    val spec' = map (pair Attrib.empty_binding) matches
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   363
    val (lthy, _, _, unfold_thms) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   364
      add_fixdefs fixes spec' lthy
40041
1f09b4c7b85e replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents: 40036
diff changeset
   365
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   366
    val blocks' = map (map fst o filter_out snd) blocks
40041
1f09b4c7b85e replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents: 40036
diff changeset
   367
    val simps : (Attrib.binding * thm) list list =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   368
      map (make_simps lthy) (unfold_thms ~~ blocks')
40041
1f09b4c7b85e replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents: 40036
diff changeset
   369
    fun mk_bind n : Attrib.binding =
45592
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 44080
diff changeset
   370
     (Binding.qualify true n (Binding.name "simps"), @{attributes [simp]})
40041
1f09b4c7b85e replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents: 40036
diff changeset
   371
    val simps1 : (Attrib.binding * thm list) list =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   372
      map (fn (n,xs) => (mk_bind n, map snd xs)) (names ~~ simps)
40041
1f09b4c7b85e replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents: 40036
diff changeset
   373
    val simps2 : (Attrib.binding * thm list) list =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   374
      map (apsnd (fn thm => [thm])) (flat simps)
40041
1f09b4c7b85e replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents: 40036
diff changeset
   375
    val (_, lthy) = lthy
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   376
      |> fold_map Local_Theory.note (simps1 @ simps2)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   377
  in
40041
1f09b4c7b85e replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents: 40036
diff changeset
   378
    lthy
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   379
  end
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   380
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 30157
diff changeset
   381
in
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   382
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   383
val add_fixrec = gen_fixrec Specification.check_spec
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   384
val add_fixrec_cmd = gen_fixrec Specification.read_spec
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 30157
diff changeset
   385
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   386
end (* local *)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   387
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   388
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   389
(*************************************************************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   390
(******************************** Parsers ********************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   391
(*************************************************************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   392
40096
4d1a8fa8cdfd change fixrec parser to not accept theorem names with (unchecked) option
huffman
parents: 40041
diff changeset
   393
val opt_thm_name' : (bool * Attrib.binding) parser =
4d1a8fa8cdfd change fixrec parser to not accept theorem names with (unchecked) option
huffman
parents: 40041
diff changeset
   394
  Parse.$$$ "(" -- Parse.$$$ "unchecked" -- Parse.$$$ ")" >> K (true, Attrib.empty_binding)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   395
    || Parse_Spec.opt_thm_name ":" >> pair false
40096
4d1a8fa8cdfd change fixrec parser to not accept theorem names with (unchecked) option
huffman
parents: 40041
diff changeset
   396
4d1a8fa8cdfd change fixrec parser to not accept theorem names with (unchecked) option
huffman
parents: 40041
diff changeset
   397
val spec' : (bool * (Attrib.binding * string)) parser =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   398
  opt_thm_name' -- Parse.prop >> (fn ((a, b), c) => (a, (b, c)))
40096
4d1a8fa8cdfd change fixrec parser to not accept theorem names with (unchecked) option
huffman
parents: 40041
diff changeset
   399
40041
1f09b4c7b85e replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents: 40036
diff changeset
   400
val alt_specs' : (bool * (Attrib.binding * string)) list parser =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   401
  let val unexpected = Scan.ahead (Parse.name || Parse.$$$ "[" || Parse.$$$ "(")
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   402
  in Parse.enum1 "|" (spec' --| Scan.option (unexpected -- Parse.!!! (Parse.$$$ "|"))) end
40041
1f09b4c7b85e replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents: 40036
diff changeset
   403
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36954
diff changeset
   404
val _ =
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36954
diff changeset
   405
  Outer_Syntax.local_theory "fixrec" "define recursive functions (HOLCF)" Keyword.thy_decl
40041
1f09b4c7b85e replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents: 40036
diff changeset
   406
    (Parse.fixes -- (Parse.where_ |-- Parse.!!! alt_specs')
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   407
      >> (fn (fixes, specs) => add_fixrec_cmd fixes specs))
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   408
33425
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   409
val setup =
37080
a2a1c8a658ef remove fixrec_simp attribute; fixrec uses default simpset from theory context instead
huffman
parents: 37079
diff changeset
   410
  Method.setup @{binding fixrec_simp}
a2a1c8a658ef remove fixrec_simp attribute; fixrec uses default simpset from theory context instead
huffman
parents: 37079
diff changeset
   411
    (Scan.succeed (SIMPLE_METHOD' o fixrec_simp_tac))
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   412
    "pattern prover for fixrec constants"
30131
6be1be402ef0 use TheoryData to keep track of pattern match combinators
huffman
parents: 29585
diff changeset
   413
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   414
end