src/HOL/Tools/function_package/fundef_prep.ML
author krauss
Wed, 08 Nov 2006 09:08:54 +0100
changeset 21240 8e75fb38522c
parent 21237 b803f9870e97
child 21255 617fdb08abe9
permissions -rw-r--r--
Made "termination by lexicographic_order" the default for "fun" definitions.
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_prep.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
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
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
Preparation step: makes auxiliary definitions and generates
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     7
proof obligations.
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     8
*)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     9
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    10
signature FUNDEF_PREP =
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    11
sig
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
    12
    val prepare_fundef : string (* defname *)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
    13
                         -> (string * typ * mixfix) (* defined symbol *)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
    14
                         -> ((string * typ) list * term list * term * term) list (* specification *)
20654
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20536
diff changeset
    15
                         -> string (* default_value, not parsed yet *)
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
    16
                         -> local_theory
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
    17
                         -> FundefCommon.prep_result * term * local_theory
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19583
diff changeset
    18
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    19
end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    20
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
    21
structure FundefPrep : FUNDEF_PREP =
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    22
struct
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
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
    25
open FundefLib
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    26
open FundefCommon
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
    27
open FundefAbbrev
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    28
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    29
(* Theory dependencies. *)
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    30
val Pair_inject = thm "Product_Type.Pair_inject";
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    31
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
    32
val acc_induct_rule = thm "FunDef.accP_induct_rule"
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    33
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    34
val ex1_implies_ex = thm "FunDef.fundef_ex1_existence"
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    35
val ex1_implies_un = thm "FunDef.fundef_ex1_uniqueness"
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    36
val ex1_implies_iff = thm "FunDef.fundef_ex1_iff"
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    37
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    38
val conjunctionI = thm "conjunctionI";
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    39
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    40
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    41
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    42
fun find_calls tree =
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    43
    let
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    44
      fun add_Ri (fixes,assumes) (_ $ arg) _ (_, xs) = ([], (fixes, assumes, arg) :: xs)
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
    45
        | add_Ri _ _ _ _ = raise Match
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
    46
    in
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    47
      rev (FundefCtxTree.traverse_tree add_Ri tree [])
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    48
    end
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    49
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    50
21188
2aa15b663cd4 minor cleanup
krauss
parents: 21100
diff changeset
    51
fun mk_compat_proof_obligations domT ranT fvar f glrs =
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    52
    let
21188
2aa15b663cd4 minor cleanup
krauss
parents: 21100
diff changeset
    53
      fun mk_impl ((qs, gs, lhs, rhs),(qs', gs', lhs', rhs')) =
2aa15b663cd4 minor cleanup
krauss
parents: 21100
diff changeset
    54
          let
2aa15b663cd4 minor cleanup
krauss
parents: 21100
diff changeset
    55
            val shift = incr_boundvars (length qs')
2aa15b663cd4 minor cleanup
krauss
parents: 21100
diff changeset
    56
          in
2aa15b663cd4 minor cleanup
krauss
parents: 21100
diff changeset
    57
            (implies $ Trueprop (eq_const domT $ shift lhs $ lhs')
2aa15b663cd4 minor cleanup
krauss
parents: 21100
diff changeset
    58
                     $ Trueprop (eq_const ranT $ shift rhs $ rhs'))
2aa15b663cd4 minor cleanup
krauss
parents: 21100
diff changeset
    59
              |> fold_rev (curry Logic.mk_implies) (map shift gs @ gs')
2aa15b663cd4 minor cleanup
krauss
parents: 21100
diff changeset
    60
              |> fold_rev (fn (n,T) => fn b => all T $ Abs(n,T,b)) (qs @ qs')
2aa15b663cd4 minor cleanup
krauss
parents: 21100
diff changeset
    61
              |> curry abstract_over fvar
2aa15b663cd4 minor cleanup
krauss
parents: 21100
diff changeset
    62
              |> curry subst_bound f
2aa15b663cd4 minor cleanup
krauss
parents: 21100
diff changeset
    63
          end
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    64
    in
21188
2aa15b663cd4 minor cleanup
krauss
parents: 21100
diff changeset
    65
      map mk_impl (unordered_pairs glrs)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    66
    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    67
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    68
21188
2aa15b663cd4 minor cleanup
krauss
parents: 21100
diff changeset
    69
fun mk_completeness (Globals {x, Pbool, ...}) clauses qglrs =
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    70
    let
21188
2aa15b663cd4 minor cleanup
krauss
parents: 21100
diff changeset
    71
        fun mk_case (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) =
2aa15b663cd4 minor cleanup
krauss
parents: 21100
diff changeset
    72
            Trueprop Pbool
2aa15b663cd4 minor cleanup
krauss
parents: 21100
diff changeset
    73
                     |> curry Logic.mk_implies (Trueprop (mk_eq (x, lhs)))
2aa15b663cd4 minor cleanup
krauss
parents: 21100
diff changeset
    74
                     |> fold_rev (curry Logic.mk_implies) gs
2aa15b663cd4 minor cleanup
krauss
parents: 21100
diff changeset
    75
                     |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    76
    in
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
    77
        Trueprop Pbool
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
    78
                 |> fold_rev (curry Logic.mk_implies o mk_case) (clauses ~~ qglrs)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
    79
                 |> mk_forall_rename ("x", x)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
    80
                 |> mk_forall_rename ("P", Pbool)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    81
    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    82
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    83
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
    84
fun mk_clause_context x ctxt (pre_qs,pre_gs,pre_lhs,pre_rhs) =
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
    85
    let
20797
c1f0bc7e7d80 renamed Variable.invent_fixes to Variable.variant_fixes;
wenzelm
parents: 20654
diff changeset
    86
      val (qs, ctxt') = Variable.variant_fixes (map fst pre_qs) ctxt
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
    87
                                           |>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    88
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
    89
      val thy = ProofContext.theory_of ctxt'
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
    90
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
    91
      fun inst t = subst_bounds (rev qs, t)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
    92
      val gs = map inst pre_gs
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
    93
      val lhs = inst pre_lhs
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
    94
      val rhs = inst pre_rhs
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
    95
                
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
    96
      val cqs = map (cterm_of thy) qs
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
    97
      val ags = map (assume o cterm_of thy) gs
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
    98
                  
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
    99
      val case_hyp = assume (cterm_of thy (Trueprop (mk_eq (x, lhs))))
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   100
    in
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   101
      ClauseContext { ctxt = ctxt', qs = qs, gs = gs, lhs = lhs, rhs = rhs, 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   102
                      cqs = cqs, ags = ags, case_hyp = case_hyp }
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   103
    end
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   104
      
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   105
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   106
(* lowlevel term function *)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   107
fun abstract_over_list vs body =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   108
  let
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   109
    exception SAME;
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   110
    fun abs lev v tm =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   111
      if v aconv tm then Bound lev
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   112
      else
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   113
        (case tm of
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   114
          Abs (a, T, t) => Abs (a, T, abs (lev + 1) v t)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   115
        | t $ u => (abs lev v t $ (abs lev v u handle SAME => u) handle SAME => t $ abs lev v u)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   116
        | _ => raise SAME);
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   117
  in 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   118
    fold_index (fn (i,v) => fn t => abs i v t handle SAME => t) vs body
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   119
  end
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   120
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   121
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   122
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   123
fun mk_clause_info globals G f no cdata qglr tree RCs GIntro_thm RIntro_thms =
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   124
    let
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   125
        val Globals {h, fvar, x, ...} = globals
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   126
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   127
        val ClauseContext { ctxt, qs, cqs, ags, ... } = cdata
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   128
        val cert = Thm.cterm_of (ProofContext.theory_of ctxt)
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   129
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   130
        (* Instantiate the GIntro thm with "f" and import into the clause context. *)
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   131
        val lGI = GIntro_thm
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   132
                    |> forall_elim (cert f)
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   133
                    |> fold forall_elim cqs
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   134
                    |> fold implies_elim_swp ags
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   135
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   136
        fun mk_call_info (rcfix, rcassm, rcarg) RI =
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   137
            let
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   138
                val llRI = RI
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   139
                             |> fold forall_elim cqs
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   140
                             |> fold (forall_elim o cert o Free) rcfix
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   141
                             |> fold implies_elim_swp ags
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   142
                             |> fold implies_elim_swp rcassm
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   143
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   144
                val h_assum =
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   145
                    Trueprop (G $ rcarg $ (h $ rcarg))
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   146
                              |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   147
                              |> fold_rev (mk_forall o Free) rcfix
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   148
                              |> Pattern.rewrite_term (ProofContext.theory_of ctxt) [(f, h)] []
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   149
                              |> abstract_over_list (rev qs)
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   150
            in
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   151
                RCInfo {RIvs=rcfix, rcarg=rcarg, CCas=rcassm, llRI=llRI, h_assum=h_assum}
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   152
            end
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   153
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   154
        val RC_infos = map2 mk_call_info RCs RIntro_thms
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   155
    in
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   156
        ClauseInfo
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   157
            {
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   158
             no=no,
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   159
             cdata=cdata,
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   160
             qglr=qglr,
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   161
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   162
             lGI=lGI, 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   163
             RCs=RC_infos,
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   164
             tree=tree
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   165
            }
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   166
    end
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   167
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   168
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   169
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   170
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   171
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   172
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   173
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   174
(* replace this by a table later*)
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   175
fun store_compat_thms 0 thms = []
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   176
  | store_compat_thms n thms =
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   177
    let
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   178
        val (thms1, thms2) = chop n thms
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   179
    in
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   180
        (thms1 :: store_compat_thms (n - 1) thms2)
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   181
    end
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   182
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   183
(* expects i <= j *)
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   184
fun lookup_compat_thm i j cts =
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   185
    nth (nth cts (i - 1)) (j - i)
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   186
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   187
(* Returns "Gsi, Gsj, lhs_i = lhs_j |-- rhs_j_f = rhs_i_f" *)
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   188
(* if j < i, then turn around *)
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   189
fun get_compat_thm thy cts i j ctxi ctxj = 
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   190
    let
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   191
        val ClauseContext {cqs=cqsi,ags=agsi,lhs=lhsi,...} = ctxi
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   192
        val ClauseContext {cqs=cqsj,ags=agsj,lhs=lhsj,...} = ctxj
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   193
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   194
        val lhsi_eq_lhsj = cterm_of thy (Trueprop (mk_eq (lhsi, lhsj)))
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   195
    in if j < i then
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   196
           let
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   197
               val compat = lookup_compat_thm j i cts
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   198
           in
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   199
               compat         (* "!!qj qi. Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   200
                |> fold forall_elim (cqsj @ cqsi) (* "Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   201
                |> fold implies_elim_swp agsj
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   202
                |> fold implies_elim_swp agsi
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   203
                |> implies_elim_swp ((assume lhsi_eq_lhsj) RS sym) (* "Gsj, Gsi, lhsi = lhsj |-- rhsj = rhsi" *)
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   204
           end
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   205
       else
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   206
           let
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   207
               val compat = lookup_compat_thm i j cts
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   208
           in
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   209
               compat        (* "!!qi qj. Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   210
                 |> fold forall_elim (cqsi @ cqsj) (* "Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   211
                 |> fold implies_elim_swp agsi
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   212
                 |> fold implies_elim_swp agsj
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   213
                 |> implies_elim_swp (assume lhsi_eq_lhsj)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   214
                 |> (fn thm => thm RS sym) (* "Gsi, Gsj, lhsi = lhsj |-- rhsj = rhsi" *)
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   215
           end
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   216
    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   217
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   218
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   219
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   220
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   221
(* Generates the replacement lemma in fully quantified form. *)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   222
fun mk_replacement_lemma thy h ih_elim clause =
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   223
    let
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   224
        val ClauseInfo {cdata=ClauseContext {qs, lhs, rhs, cqs, ags, case_hyp, ...}, RCs, tree, ...} = clause
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   225
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   226
        val ih_elim_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ih_elim
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   227
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   228
        val Ris = map (fn RCInfo {llRI, ...} => llRI) RCs
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   229
        val h_assums = map (fn RCInfo {h_assum, ...} => assume (cterm_of thy (subst_bounds (rev qs, h_assum)))) RCs
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   230
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   231
        val ih_elim_case_inst = instantiate' [] [NONE, SOME (cterm_of thy h)] ih_elim_case (* Should be done globally *)
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   232
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   233
        val (eql, _) = FundefCtxTree.rewrite_by_tree thy h ih_elim_case_inst (Ris ~~ h_assums) tree
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   234
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   235
        val replace_lemma = (eql RS meta_eq_to_obj_eq)
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   236
                                |> implies_intr (cprop_of case_hyp)
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   237
                                |> fold_rev (implies_intr o cprop_of) h_assums
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   238
                                |> fold_rev (implies_intr o cprop_of) ags
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   239
                                |> fold_rev forall_intr cqs
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   240
    in
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   241
      replace_lemma
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   242
    end
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   243
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   244
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   245
fun mk_uniqueness_clause thy globals f compat_store clausei clausej RLj =
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   246
    let
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   247
        val Globals {h, y, x, fvar, ...} = globals
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   248
        val ClauseInfo {no=i, cdata=cctxi as ClauseContext {ctxt=ctxti, lhs=lhsi, case_hyp, ...}, ...} = clausei
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   249
        val ClauseInfo {no=j, qglr=cdescj, RCs=RCsj, ...} = clausej
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   250
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   251
        val cctxj as ClauseContext {ags = agsj', lhs = lhsj', rhs = rhsj', qs = qsj', cqs = cqsj', ...} 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   252
            = mk_clause_context x ctxti cdescj
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   253
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   254
        val rhsj'h = Pattern.rewrite_term thy [(fvar,h)] [] rhsj'
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   255
        val compat = get_compat_thm thy compat_store i j cctxi cctxj
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   256
        val Ghsj' = map (fn RCInfo {h_assum, ...} => assume (cterm_of thy (subst_bounds (rev qsj', h_assum)))) RCsj
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   257
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   258
        val RLj_import = 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   259
            RLj |> fold forall_elim cqsj'
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   260
                |> fold implies_elim_swp agsj'
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   261
                |> fold implies_elim_swp Ghsj'
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   262
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   263
        val y_eq_rhsj'h = assume (cterm_of thy (Trueprop (mk_eq (y, rhsj'h))))
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   264
        val lhsi_eq_lhsj' = assume (cterm_of thy (Trueprop (mk_eq (lhsi, lhsj')))) (* lhs_i = lhs_j' |-- lhs_i = lhs_j' *)
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   265
    in
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   266
        (trans OF [case_hyp, lhsi_eq_lhsj']) (* lhs_i = lhs_j' |-- x = lhs_j' *)
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   267
        |> implies_elim RLj_import (* Rj1' ... Rjk', lhs_i = lhs_j' |-- rhs_j'_h = rhs_j'_f *)
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   268
        |> (fn it => trans OF [it, compat]) (* lhs_i = lhs_j', Gj', Rj1' ... Rjk' |-- rhs_j'_h = rhs_i_f *)
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   269
        |> (fn it => trans OF [y_eq_rhsj'h, it]) (* lhs_i = lhs_j', Gj', Rj1' ... Rjk', y = rhs_j_h' |-- y = rhs_i_f *)
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   270
        |> fold_rev (implies_intr o cprop_of) Ghsj'
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   271
        |> fold_rev (implies_intr o cprop_of) agsj' (* lhs_i = lhs_j' , y = rhs_j_h' |-- Gj', Rj1'...Rjk' ==> y = rhs_i_f *)
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   272
        |> implies_intr (cprop_of y_eq_rhsj'h)
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   273
        |> implies_intr (cprop_of lhsi_eq_lhsj')
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   274
        |> fold_rev forall_intr (cterm_of thy h :: cqsj')
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   275
    end
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   276
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   277
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   278
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   279
fun mk_uniqueness_case thy globals G f ihyp ih_intro G_cases compat_store clauses rep_lemmas clausei =
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   280
    let
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   281
        val Globals {x, y, ranT, fvar, ...} = globals
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   282
        val ClauseInfo {cdata = ClauseContext {lhs, rhs, qs, cqs, ags, case_hyp, ...}, lGI, RCs, ...} = clausei
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   283
        val rhsC = Pattern.rewrite_term thy [(fvar, f)] [] rhs
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   284
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   285
        val ih_intro_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ih_intro
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   286
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   287
        fun prep_RC (RCInfo {llRI, RIvs, CCas, ...}) = (llRI RS ih_intro_case)
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   288
                                                            |> fold_rev (implies_intr o cprop_of) CCas
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   289
                                                            |> fold_rev (forall_intr o cterm_of thy o Free) RIvs
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   290
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   291
        val existence = fold (curry op COMP o prep_RC) RCs lGI
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   292
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   293
        val P = cterm_of thy (mk_eq (y, rhsC))
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   294
        val G_lhs_y = assume (cterm_of thy (Trueprop (G $ lhs $ y)))
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   295
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   296
        val unique_clauses = map2 (mk_uniqueness_clause thy globals f compat_store clausei) clauses rep_lemmas
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   297
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   298
        val uniqueness = G_cases
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   299
                           |> forall_elim (cterm_of thy lhs)
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   300
                           |> forall_elim (cterm_of thy y)
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   301
                           |> forall_elim P
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   302
                           |> implies_elim_swp G_lhs_y
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   303
                           |> fold implies_elim_swp unique_clauses
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   304
                           |> implies_intr (cprop_of G_lhs_y)
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   305
                           |> forall_intr (cterm_of thy y)
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   306
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   307
        val P2 = cterm_of thy (lambda y (G $ lhs $ y)) (* P2 y := (lhs, y): G *)
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   308
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   309
        val exactly_one =
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   310
            ex1I |> instantiate' [SOME (ctyp_of thy ranT)] [SOME P2, SOME (cterm_of thy rhsC)]
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   311
                 |> curry (op COMP) existence
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   312
                 |> curry (op COMP) uniqueness
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   313
                 |> simplify (HOL_basic_ss addsimps [case_hyp RS sym])
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   314
                 |> implies_intr (cprop_of case_hyp)
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   315
                 |> fold_rev (implies_intr o cprop_of) ags
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   316
                 |> fold_rev forall_intr cqs
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   317
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   318
        val function_value =
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   319
            existence
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   320
              |> implies_intr ihyp
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   321
              |> implies_intr (cprop_of case_hyp)
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   322
              |> forall_intr (cterm_of thy x)
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   323
              |> forall_elim (cterm_of thy lhs)
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   324
              |> curry (op RS) refl
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   325
    in
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   326
        (exactly_one, function_value)
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   327
    end
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   328
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   329
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   330
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   331
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   332
fun prove_stuff thy congs globals G f R clauses complete compat compat_store G_elim f_def =
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   333
    let
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   334
        val Globals {h, domT, ranT, x, ...} = globals
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   335
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   336
        val inst_ex1_ex =  f_def RS ex1_implies_ex
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   337
        val inst_ex1_un =  f_def RS ex1_implies_un
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   338
        val inst_ex1_iff = f_def RS ex1_implies_iff
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   339
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   340
        (* Inductive Hypothesis: !!z. (z,x):R ==> EX!y. (z,y):G *)
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   341
        val ihyp = all domT $ Abs ("z", domT,
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   342
                                   implies $ Trueprop (R $ Bound 0 $ x)
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   343
                                           $ Trueprop (Const ("Ex1", (ranT --> boolT) --> boolT) $
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   344
                                                             Abs ("y", ranT, G $ Bound 1 $ Bound 0)))
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   345
                       |> cterm_of thy
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   346
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   347
        val ihyp_thm = assume ihyp |> forall_elim_vars 0
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   348
        val ih_intro = ihyp_thm RS inst_ex1_ex
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   349
        val ih_elim = ihyp_thm RS inst_ex1_un
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   350
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   351
        val _ = Output.debug "Proving Replacement lemmas..."
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   352
        val repLemmas = map (mk_replacement_lemma thy h ih_elim) clauses
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   353
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   354
        val _ = Output.debug "Proving cases for unique existence..."
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   355
        val (ex1s, values) = 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   356
            split_list (map (mk_uniqueness_case thy globals G f ihyp ih_intro G_elim compat_store clauses repLemmas) clauses)
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   357
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   358
        val _ = Output.debug "Proving: Graph is a function" (* FIXME: Rewrite this proof. *)
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   359
        val graph_is_function = complete
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   360
                                  |> forall_elim_vars 0
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   361
                                  |> fold (curry op COMP) ex1s
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   362
                                  |> implies_intr (ihyp)
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   363
                                  |> implies_intr (cterm_of thy (Trueprop (mk_acc domT R $ x)))
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   364
                                  |> forall_intr (cterm_of thy x)
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   365
                                  |> (fn it => Drule.compose_single (it, 2, acc_induct_rule)) (* "EX! y. (?x,y):G" *)
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   366
                                  |> (fn it => fold (forall_intr o cterm_of thy) (term_vars (prop_of it)) it)
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   367
                                  |> Drule.close_derivation
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   368
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   369
        val goal = complete COMP (graph_is_function COMP conjunctionI)
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   370
                            |> Drule.close_derivation
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   371
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   372
        val goalI = Goal.protect goal
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   373
                                 |> fold_rev (implies_intr o cprop_of) compat
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   374
                                 |> implies_intr (cprop_of complete)
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   375
    in
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   376
      (prop_of goal, goalI, inst_ex1_iff, values)
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   377
    end
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   378
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   379
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   380
fun define_graph Gname fvar domT ranT clauses RCss lthy =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   381
    let
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   382
      val GT = domT --> ranT --> boolT
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   383
      val Gvar = Free (the_single (Variable.variant_frees lthy [] [(Gname, GT)]))
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   384
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   385
      fun mk_GIntro (ClauseContext {qs, gs, lhs, rhs, ...}) RCs =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   386
          let
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   387
            fun mk_h_assm (rcfix, rcassm, rcarg) =
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   388
                Trueprop (Gvar $ rcarg $ (fvar $ rcarg))
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   389
                          |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   390
                          |> fold_rev (mk_forall o Free) rcfix
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   391
          in
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   392
            Trueprop (Gvar $ lhs $ rhs)
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   393
                      |> fold_rev (curry Logic.mk_implies o mk_h_assm) RCs
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   394
                      |> fold_rev (curry Logic.mk_implies) gs
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   395
                      |> fold_rev mk_forall (fvar :: qs)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   396
          end
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   397
          
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   398
      val G_intros = map2 mk_GIntro clauses RCss
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   399
                     
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   400
      val (GIntro_thms, (G, G_elim, lthy)) = 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   401
          FundefInductiveWrap.inductive_def G_intros ((dest_Free Gvar, NoSyn), lthy)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   402
    in
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   403
      ((G, GIntro_thms, G_elim), lthy)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   404
    end
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   405
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   406
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   407
20654
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20536
diff changeset
   408
fun define_function fdefname (fname, mixfix) domT ranT G default lthy =
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   409
    let
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   410
      val f_def = 
20654
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20536
diff changeset
   411
          Abs ("x", domT, Const ("FunDef.THE_default", ranT --> (ranT --> boolT) --> ranT) $ (default $ Bound 0) $
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   412
                                Abs ("y", ranT, G $ Bound 1 $ Bound 0))
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   413
              |> Envir.beta_norm (* Fixme: LocalTheory.def does not work if not beta-normal *)
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   414
          
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   415
      val ((f, (_, f_defthm)), lthy) = LocalTheory.def ((fname ^ "C", mixfix), ((fdefname, []), f_def)) lthy
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   416
    in
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   417
      ((f, f_defthm), lthy)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   418
    end
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   419
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   420
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   421
fun define_recursion_relation Rname domT ranT fvar f qglrs clauses RCss lthy =
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19583
diff changeset
   422
    let
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   423
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   424
      val RT = domT --> domT --> boolT
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   425
      val Rvar = Free (the_single (Variable.variant_frees lthy [] [(Rname, RT)]))
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   426
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   427
      fun mk_RIntro (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) (rcfix, rcassm, rcarg) =
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   428
          Trueprop (Rvar $ rcarg $ lhs)
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   429
                    |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   430
                    |> fold_rev (curry Logic.mk_implies) gs
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   431
                    |> fold_rev (mk_forall o Free) rcfix
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   432
                    |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   433
                    (* "!!qs xs. CS ==> G => (r, lhs) : R" *)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   434
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   435
      val R_intross = map2 (map o mk_RIntro) (clauses ~~ qglrs) RCss
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   436
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   437
      val (RIntro_thmss, (R, R_elim, lthy)) = 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   438
          fold_burrow FundefInductiveWrap.inductive_def R_intross ((dest_Free Rvar, NoSyn), lthy)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   439
    in
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   440
      ((R, RIntro_thmss, R_elim), lthy)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   441
    end
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19583
diff changeset
   442
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   443
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   444
fun fix_globals domT ranT fvar ctxt =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   445
    let
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   446
      val ([h, y, x, z, a, D, P, Pbool],ctxt') = 
20797
c1f0bc7e7d80 renamed Variable.invent_fixes to Variable.variant_fixes;
wenzelm
parents: 20654
diff changeset
   447
          Variable.variant_fixes ["h_fd", "y_fd", "x_fd", "z_fd", "a_fd", "D_fd", "P_fd", "Pb_fd"] ctxt
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   448
    in
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   449
      (Globals {h = Free (h, domT --> ranT),
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   450
                y = Free (y, ranT),
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   451
                x = Free (x, domT),
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   452
                z = Free (z, domT),
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   453
                a = Free (a, domT),
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   454
                D = Free (D, domT --> boolT),
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   455
                P = Free (P, domT --> boolT),
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   456
                Pbool = Free (Pbool, boolT),
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   457
                fvar = fvar,
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   458
                domT = domT,
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   459
                ranT = ranT
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   460
               },
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   461
       ctxt')
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   462
    end
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   463
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   464
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   465
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   466
fun inst_RC thy fvar f (rcfix, rcassm, rcarg) =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   467
    let
21237
b803f9870e97 untabified
krauss
parents: 21188
diff changeset
   468
      fun inst_term t = subst_bound(f, abstract_over (fvar, t))
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   469
    in
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   470
      (rcfix, map (assume o cterm_of thy o inst_term o prop_of) rcassm, inst_term rcarg)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   471
    end
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   472
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   473
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   474
20654
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20536
diff changeset
   475
fun prepare_fundef defname (fname, fT, mixfix) abstract_qglrs default_str lthy =
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   476
    let
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   477
      val fvar = Free (fname, fT)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   478
      val domT = domain_type fT
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   479
      val ranT = range_type fT
20654
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20536
diff changeset
   480
                            
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20536
diff changeset
   481
      val [default] = fst (Variable.importT_terms (fst (ProofContext.read_termTs lthy (K false) (K NONE) (K NONE) [] [(default_str, fT)])) lthy) (* FIXME *)
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   482
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   483
      val congs = get_fundef_congs (Context.Proof lthy)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   484
      val (globals, ctxt') = fix_globals domT ranT fvar lthy
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   485
21100
cda93bbf35db Fixed bug in the handling of congruence rules
krauss
parents: 21051
diff changeset
   486
      val Globals { x, h, ... } = globals
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   487
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   488
      val clauses = map (mk_clause_context x ctxt') abstract_qglrs 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   489
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   490
      val n = length abstract_qglrs
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   491
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   492
      val congs_deps = map (fn c => (c, FundefCtxTree.cong_deps c)) (congs @ FundefCtxTree.add_congs) (* FIXME: Save in theory *)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   493
              
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   494
      fun build_tree (ClauseContext { ctxt, rhs, ...}) = 
21100
cda93bbf35db Fixed bug in the handling of congruence rules
krauss
parents: 21051
diff changeset
   495
            FundefCtxTree.mk_tree congs_deps (fname, fT) h ctxt rhs
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   496
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   497
      val trees = map build_tree clauses
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   498
      val RCss = map find_calls trees
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   499
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   500
      val ((G, GIntro_thms, G_elim), lthy) = define_graph (defname ^ "_graph") fvar domT ranT clauses RCss lthy
20654
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20536
diff changeset
   501
      val ((f, f_defthm), lthy) = define_function (defname ^ "_sum_def") (fname, mixfix) domT ranT G default lthy
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   502
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   503
      val RCss = map (map (inst_RC (ProofContext.theory_of lthy) fvar f)) RCss
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   504
      val trees = map (FundefCtxTree.inst_tree (ProofContext.theory_of lthy) fvar f) trees
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19583
diff changeset
   505
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   506
      val ((R, RIntro_thmss, R_elim), lthy) = 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   507
          define_recursion_relation (defname ^ "_rel") domT ranT fvar f abstract_qglrs clauses RCss lthy
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19876
diff changeset
   508
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20797
diff changeset
   509
      val dom_abbrev = Logic.mk_equals (Free (defname ^ "_dom", domT --> boolT), mk_acc domT R)
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   510
      val lthy = Specification.abbreviation_i ("", false) [(NONE, dom_abbrev)] lthy
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19583
diff changeset
   511
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   512
      val newthy = ProofContext.theory_of lthy
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   513
      val clauses = map (transfer_clause_ctx newthy) clauses
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   514
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   515
      val cert = cterm_of (ProofContext.theory_of lthy)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   516
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   517
      val xclauses = map7 (mk_clause_info globals G f) (1 upto n) clauses abstract_qglrs trees RCss GIntro_thms RIntro_thmss
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   518
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   519
      val complete = mk_completeness globals clauses abstract_qglrs |> cert |> assume
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   520
      val compat = mk_compat_proof_obligations domT ranT fvar f abstract_qglrs |> map (cert #> assume)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   521
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   522
      val compat_store = store_compat_thms n compat
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19922
diff changeset
   523
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   524
      val (goal, goalI, ex1_iff, values) = prove_stuff newthy congs globals G f R xclauses complete compat compat_store G_elim f_defthm
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19583
diff changeset
   525
    in
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   526
        (Prep {globals = globals, f = f, G = G, R = R, goal = goal, goalI = goalI, values = values, clauses = xclauses, ex1_iff = ex1_iff, R_cases = R_elim}, f,
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20071
diff changeset
   527
         lthy)
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19583
diff changeset
   528
    end
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19583
diff changeset
   529
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19583
diff changeset
   530
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19583
diff changeset
   531
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19583
diff changeset
   532
19876
wenzelm
parents: 19781
diff changeset
   533
end