src/HOL/Tools/function_package/fundef_datatype.ML
author wenzelm
Thu, 12 Oct 2006 22:57:15 +0200
changeset 20999 9131d8e96dba
parent 20875 95d24e8d117e
child 21051 c49467a9c1e1
permissions -rw-r--r--
Toplevel.local_theory: proper target; removed dead code;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     1
(*  Title:      HOL/Tools/function_package/fundef_datatype.ML
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     2
    ID:         $Id$
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     3
    Author:     Alexander Krauss, TU Muenchen
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     4
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
     5
A package for general recursive function definitions.
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     6
A tactic to prove completeness of datatype patterns.
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     7
*)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     8
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
     9
signature FUNDEF_DATATYPE =
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    10
sig
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    11
    val pat_complete_tac: int -> tactic
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    12
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
    13
    val pat_completeness : method
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    14
    val setup : theory -> theory
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    15
end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    16
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    17
structure FundefDatatype : FUNDEF_DATATYPE =
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    18
struct
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    19
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19564
diff changeset
    20
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    21
fun mk_argvar i T = Free ("_av" ^ (string_of_int i), T)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    22
fun mk_patvar i T = Free ("_pv" ^ (string_of_int i), T)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    23
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    24
fun inst_free var inst thm =
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    25
    forall_elim inst (forall_intr var thm)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    26
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    27
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    28
fun inst_case_thm thy x P thm =
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    29
    let
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    30
        val [Pv, xv] = term_vars (prop_of thm)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    31
    in
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    32
        cterm_instantiate [(cterm_of thy xv, cterm_of thy x), (cterm_of thy Pv, cterm_of thy P)] thm
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    33
    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    34
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    35
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    36
fun invent_vars constr i =
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    37
    let
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    38
        val Ts = binder_types (fastype_of constr)
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    39
        val j = i + length Ts
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    40
        val is = i upto (j - 1)
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    41
        val avs = map2 mk_argvar is Ts
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    42
        val pvs = map2 mk_patvar is Ts
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    43
    in
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    44
        (avs, pvs, j)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    45
    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    46
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    47
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    48
fun filter_pats thy cons pvars [] = []
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    49
  | filter_pats thy cons pvars (([], thm) :: pts) = raise Match
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    50
  | filter_pats thy cons pvars ((pat :: pats, thm) :: pts) =
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    51
    case pat of
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    52
        Free _ => let val inst = list_comb (cons, pvars)
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    53
                 in (inst :: pats, inst_free (cterm_of thy pat) (cterm_of thy inst) thm)
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    54
                    :: (filter_pats thy cons pvars pts) end
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    55
      | _ => if fst (strip_comb pat) = cons
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    56
             then (pat :: pats, thm) :: (filter_pats thy cons pvars pts)
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    57
             else filter_pats thy cons pvars pts
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    58
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    59
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    60
fun inst_constrs_of thy (T as Type (name, _)) =
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    61
        map (fn (Cn,CT) => Envir.subst_TVars (Sign.typ_match thy (body_type CT, T) Vartab.empty) (Const (Cn, CT)))
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    62
            (the (DatatypePackage.get_datatype_constrs thy name))
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    63
  | inst_constrs_of thy _ = raise Match
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    64
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    65
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    66
fun transform_pat thy avars c_assum ([] , thm) = raise Match
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    67
  | transform_pat thy avars c_assum (pat :: pats, thm) =
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    68
    let
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    69
        val (_, subps) = strip_comb pat
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    70
        val eqs = map (cterm_of thy o HOLogic.mk_Trueprop o HOLogic.mk_eq) (avars ~~ subps)
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    71
        val a_eqs = map assume eqs
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    72
        val c_eq_pat = simplify (HOL_basic_ss addsimps a_eqs) c_assum
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    73
    in
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    74
        (subps @ pats, fold_rev implies_intr eqs
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    75
                                (implies_elim thm c_eq_pat))
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    76
    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    77
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    78
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    79
exception COMPLETENESS
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    80
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    81
fun constr_case thy P idx (v :: vs) pats cons =
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    82
    let
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    83
        val (avars, pvars, newidx) = invent_vars cons idx
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    84
        val c_hyp = cterm_of thy (HOLogic.mk_Trueprop (HOLogic.mk_eq (v, list_comb (cons, avars))))
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    85
        val c_assum = assume c_hyp
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    86
        val newpats = map (transform_pat thy avars c_assum) (filter_pats thy cons pvars pats)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    87
    in
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    88
        o_alg thy P newidx (avars @ vs) newpats
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    89
              |> implies_intr c_hyp
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    90
              |> fold_rev (forall_intr o cterm_of thy) avars
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    91
    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    92
  | constr_case _ _ _ _ _ _ = raise Match
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    93
and o_alg thy P idx [] (([], Pthm) :: _)  = Pthm
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    94
  | o_alg thy P idx (v :: vs) [] = raise COMPLETENESS
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    95
  | o_alg thy P idx (v :: vs) pts =
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    96
    if forall (is_Free o hd o fst) pts (* Var case *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    97
    then o_alg thy P idx vs (map (fn (pv :: pats, thm) =>
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
    98
                               (pats, refl RS (inst_free (cterm_of thy pv) (cterm_of thy v) thm))) pts)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    99
    else (* Cons case *)
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   100
         let
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   101
             val T = fastype_of v
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   102
             val (tname, _) = dest_Type T
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   103
             val {exhaustion=case_thm, ...} = DatatypePackage.the_datatype thy tname
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   104
             val constrs = inst_constrs_of thy T
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   105
             val c_cases = map (constr_case thy P idx (v :: vs) pts) constrs
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   106
         in
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   107
             inst_case_thm thy v P case_thm
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   108
                           |> fold (curry op COMP) c_cases
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   109
         end
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   110
  | o_alg _ _ _ _ _ = raise Match
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   111
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   112
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   113
fun prove_completeness thy x P qss pats =
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   114
    let
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   115
        fun mk_assum qs pat = Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (x,pat)),
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   116
                                                HOLogic.mk_Trueprop P)
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   117
                                               |> fold_rev mk_forall qs
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   118
                                               |> cterm_of thy
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   119
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   120
        val hyps = map2 mk_assum qss pats
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   121
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   122
        fun inst_hyps hyp qs = fold (forall_elim o cterm_of thy) qs (assume hyp)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   123
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   124
        val assums = map2 inst_hyps hyps qss
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   125
    in
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   126
        o_alg thy P 2 [x] (map2 (pair o single) pats assums)
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   127
              |> fold_rev implies_intr hyps
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   128
    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   129
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   130
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   131
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   132
fun pat_complete_tac i thm =
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   133
    let
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19770
diff changeset
   134
      val thy = theory_of_thm thm
984ae977f7aa Fixed name clash.
krauss
parents: 19770
diff changeset
   135
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   136
        val subgoal = nth (prems_of thm) (i - 1)   (* FIXME SUBGOAL tactical *)
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19770
diff changeset
   137
984ae977f7aa Fixed name clash.
krauss
parents: 19770
diff changeset
   138
        val ([P, x], subgf) = dest_all_all subgoal
984ae977f7aa Fixed name clash.
krauss
parents: 19770
diff changeset
   139
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   140
        val assums = Logic.strip_imp_prems subgf
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   141
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   142
        fun pat_of assum =
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   143
            let
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   144
                val (qs, imp) = dest_all_all assum
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   145
            in
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   146
                case Logic.dest_implies imp of
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   147
                    (_ $ (_ $ _ $ pat), _) => (qs, pat)
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   148
                  | _ => raise COMPLETENESS
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   149
            end
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   150
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   151
        val (qss, pats) = split_list (map pat_of assums)
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   152
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   153
        val complete_thm = prove_completeness thy x P qss pats
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19770
diff changeset
   154
                                              |> forall_intr (cterm_of thy x)
984ae977f7aa Fixed name clash.
krauss
parents: 19770
diff changeset
   155
                                              |> forall_intr (cterm_of thy P)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   156
    in
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   157
        Seq.single (Drule.compose_single(complete_thm, i, thm))
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   158
    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   159
    handle COMPLETENESS => Seq.empty
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   160
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   161
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   162
val pat_completeness = Method.SIMPLE_METHOD (pat_complete_tac 1)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   163
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   164
val by_pat_completeness_simp =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   165
    Proof.global_terminal_proof
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   166
      (Method.Basic (K pat_completeness),
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   167
       SOME (Method.Source (Args.src (("simp_all", []), Position.none))))
20999
9131d8e96dba Toplevel.local_theory: proper target;
wenzelm
parents: 20875
diff changeset
   168
         (* FIXME avoid dynamic scoping of method name! *)
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   169
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   170
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   171
val setup =
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   172
    Method.add_methods [("pat_completeness", Method.no_args pat_completeness, "Completeness prover for datatype patterns")]
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   173
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   174
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   175
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   176
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   177
local structure P = OuterParse and K = OuterKeyword in
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   178
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   179
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   180
fun or_list1 s = P.enum1 "|" s
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   181
val otherwise = P.$$$ "(" |-- P.$$$ "otherwise" --| P.$$$ ")"
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   182
val statement_ow = P.and_list1 (P.opt_thm_name ":" -- Scan.repeat1 (P.prop -- Scan.optional (otherwise >> K true) false))
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   183
val statements_ow = or_list1 statement_ow
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   184
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   185
val funP =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   186
  OuterSyntax.command "fun" "define general recursive functions (short version)" K.thy_decl
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   187
  ((P.opt_locale_target -- P.fixes --| P.$$$ "where" -- statements_ow)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   188
     >> (fn ((target, fixes), statements) =>
20999
9131d8e96dba Toplevel.local_theory: proper target;
wenzelm
parents: 20875
diff changeset
   189
            Toplevel.print o
9131d8e96dba Toplevel.local_theory: proper target;
wenzelm
parents: 20875
diff changeset
   190
            Toplevel.local_theory target
9131d8e96dba Toplevel.local_theory: proper target;
wenzelm
parents: 20875
diff changeset
   191
              (FundefPackage.add_fundef fixes statements FundefCommon.fun_config
9131d8e96dba Toplevel.local_theory: proper target;
wenzelm
parents: 20875
diff changeset
   192
               #> by_pat_completeness_simp)));
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   193
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   194
val _ = OuterSyntax.add_parsers [funP];
20344
d02b43ea722e avoid low-level tsig;
wenzelm
parents: 19922
diff changeset
   195
end
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20344
diff changeset
   196
20875
95d24e8d117e TheoryTarget.init;
wenzelm
parents: 20654
diff changeset
   197
end