src/Pure/Isar/local_defs.ML
author wenzelm
Sat, 28 Jan 2006 17:29:06 +0100
changeset 18830 34b51dcdc570
child 18840 ce16e2bad548
permissions -rw-r--r--
Basic operations on local definitions.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/local_defs.ML
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
     3
    Author:     Makarius
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
     4
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
     5
Basic operations on local definitions.
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
     6
*)
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
     7
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
     8
signature LOCAL_DEFS =
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
     9
sig
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    10
  val mk_def: ProofContext.context -> (string * term) list -> term list
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    11
  val cert_def: ProofContext.context -> term -> string * term
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    12
  val abs_def: term -> (string * typ) * term
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    13
  val derived_def: ProofContext.context -> term ->
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    14
    ((string * typ) * term) * (ProofContext.context -> term -> thm -> thm)
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    15
  val def_export: ProofContext.export
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    16
  val add_def: string * term -> ProofContext.context ->
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    17
    ((string * typ) * thm) * ProofContext.context
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    18
end;
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    19
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    20
structure LocalDefs: LOCAL_DEFS =
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    21
struct
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    22
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    23
(* prepare defs *)
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    24
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    25
fun mk_def ctxt args =
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    26
  let
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    27
    val (xs, rhss) = split_list args;
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    28
    val (bind, _) = ProofContext.bind_fixes xs ctxt;
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    29
    val lhss = map (fn (x, rhs) => bind (Free (x, Term.fastype_of rhs))) args;
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    30
  in map Logic.mk_equals (lhss ~~ rhss) end;
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    31
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    32
fun cert_def ctxt eq =
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    33
  let
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    34
    fun err msg = cat_error msg
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    35
      ("The error(s) above occurred in definition: " ^ ProofContext.string_of_term ctxt eq);
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    36
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    37
    val (lhs, rhs) = Logic.dest_equals (Term.strip_all_body eq)
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    38
      handle TERM _ => err "Not a meta-equality (==)";
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    39
    val (f, xs) = Term.strip_comb (Pattern.beta_eta_contract lhs);
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    40
    val (c, _) = Term.dest_Free f handle TERM _ =>
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    41
      err "Head of lhs must be a free/fixed variable";
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    42
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    43
    fun check_arg (Bound _) = true
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    44
      | check_arg (Free (x, _)) = not (ProofContext.is_fixed ctxt x)
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    45
      | check_arg t = (case try Logic.dest_type t of SOME (TFree _) => true | _ => false);
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    46
    fun close_arg (Free (x, T)) t = Term.all T $ Term.absfree (x, T, t)
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    47
      | close_arg (Const ("TYPE", T)) t = Term.all T $ Term.absdummy (T, t)
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    48
      | close_arg _ t = t;
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    49
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    50
    val extra_frees = Term.fold_aterms (fn v as Free (x, _) =>
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    51
      if ProofContext.is_fixed ctxt x orelse member (op aconv) xs v then I
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    52
      else insert (op =) x | _ => I) rhs [];
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    53
  in
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    54
    if not (forall check_arg xs) orelse has_duplicates (op =) xs then
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    55
      err "Arguments of lhs must be distinct free/bound variables"
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    56
    else if not (null extra_frees) then
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    57
      err ("Extra free variables on rhs: " ^ commas_quote extra_frees)
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    58
    else if Term.exists_subterm (fn t => t = f) rhs then
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    59
      err "Element to be defined occurs on rhs"
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    60
    else (c, fold_rev close_arg xs eq)
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    61
  end;
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    62
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    63
fun abs_def eq =
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    64
  let
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    65
    val body = Term.strip_all_body eq;
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    66
    val vars = map Free (Term.rename_wrt_term body (Term.strip_all_vars eq));
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    67
    val (lhs, rhs) = Logic.dest_equals (Term.subst_bounds (vars, body));
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    68
    val (f, xs) = Term.strip_comb (Pattern.beta_eta_contract lhs);
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    69
    val eq' = Term.list_abs_free (map Term.dest_Free xs, rhs);
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    70
  in (Term.dest_Free f, eq') end;
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    71
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    72
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    73
(* derived defs -- potentially within the object-logic *)
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    74
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    75
fun derived_def ctxt prop =
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    76
  let
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    77
    val thy = ProofContext.theory_of ctxt;
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    78
    val ((c, T), rhs) = prop
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    79
      |> Thm.cterm_of thy
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    80
      |> ObjectLogic.meta_rewrite_cterm
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    81
      |> (snd o Logic.dest_equals o Thm.prop_of)
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    82
      |> Logic.strip_imp_concl
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    83
      |> (snd o cert_def ctxt)
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    84
      |> abs_def;
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    85
    fun prove ctxt' t def =
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    86
      let
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    87
        val thy' = ProofContext.theory_of ctxt';
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    88
        val prop' = Term.subst_atomic [(Free (c, T), t)] prop;
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    89
        val frees = Term.fold_aterms (fn Free (x, _) =>
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    90
          if ProofContext.is_fixed ctxt' x then I else insert (op =) x | _ => I) prop' [];
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    91
      in
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    92
        Goal.prove thy' frees [] prop' (K (ALLGOALS
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    93
          (ObjectLogic.meta_rewrite_tac THEN'
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    94
            Tactic.rewrite_goal_tac [def] THEN'
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    95
            Tactic.resolve_tac [Drule.reflexive_thm])))
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    96
        handle ERROR msg => cat_error msg "Failed to prove definitional specification."
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    97
      end;
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    98
  in (((c, T), rhs), prove) end;
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    99
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   100
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   101
(* export defs *)
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   102
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   103
fun head_of_def cprop =
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   104
  #1 (Term.strip_comb (#1 (Logic.dest_equals (Term.strip_all_body (Thm.term_of cprop)))))
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   105
  |> Thm.cterm_of (Thm.theory_of_cterm cprop);
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   106
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   107
fun def_export _ cprops thm =
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   108
  thm
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   109
  |> Drule.implies_intr_list cprops
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   110
  |> Drule.forall_intr_list (map head_of_def cprops)
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   111
  |> Drule.forall_elim_vars 0
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   112
  |> RANGE (replicate (length cprops) (Tactic.rtac Drule.reflexive_thm)) 1;
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   113
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   114
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   115
(* add defs *)
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   116
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   117
fun add_def (x, t) ctxt =
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   118
  let
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   119
    val [eq] = mk_def ctxt [(x, t)];
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   120
    val x' = Term.dest_Free (fst (Logic.dest_equals eq));
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   121
  in
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   122
    ctxt
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   123
    |> ProofContext.add_fixes_i [(x, NONE, NoSyn)] |> snd
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   124
    |> ProofContext.add_assms_i def_export [(("", []), [(eq, ([], []))])]
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   125
    |>> (fn [(_, [th])] => (x', th))
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   126
  end;
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   127
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   128
end;