src/HOL/typedef.ML
author paulson
Thu, 21 Mar 1996 13:02:26 +0100
changeset 1601 0ef6ea27ab15
parent 1475 7f5a4cd08209
child 2238 c72a23bbe762
permissions -rw-r--r--
Changes required by removal of the theory argument of Theorem
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1264
diff changeset
     1
(*  Title:      HOL/typedef.ML
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     3
    Author:     Markus Wenzel, TU Muenchen
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1264
diff changeset
     5
Internal interface for typedef definitions.
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     6
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     7
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1264
diff changeset
     8
signature TYPEDEF =
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     9
sig
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
  val prove_nonempty: cterm -> thm list -> tactic option -> thm
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1264
diff changeset
    11
  val add_typedef: string -> string * string list * mixfix ->
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    12
    string -> string list -> thm list -> tactic option -> theory -> theory
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1264
diff changeset
    13
  val add_typedef_i: string -> string * string list * mixfix ->
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    14
    term -> string list -> thm list -> tactic option -> theory -> theory
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    15
end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    16
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1264
diff changeset
    17
structure Typedef: TYPEDEF =
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    18
struct
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    19
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    20
open Syntax Logic HOLogic;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    21
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    22
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    23
(* prove non-emptyness of a set *)   (*exception ERROR*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    25
val is_def = is_equals o #prop o rep_thm;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    26
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    27
fun prove_nonempty cset thms usr_tac =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    28
  let
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    29
    val {T = setT, t = set, maxidx, sign} = rep_cterm cset;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    30
    val T = dest_setT setT;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    31
    val goal =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    32
      cterm_of sign (mk_Trueprop (mk_mem (Var (("x", maxidx + 1), T), set)));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    33
    val tac =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    34
      TRY (rewrite_goals_tac (filter is_def thms)) THEN
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    35
      TRY (REPEAT_FIRST (resolve_tac (filter_out is_def thms))) THEN
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    36
      if_none usr_tac (TRY (ALLGOALS (fast_tac set_cs)));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    37
  in
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    38
    prove_goalw_cterm [] goal (K [tac])
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    39
  end
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    40
  handle ERROR =>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    41
    error ("Failed to prove non-emptyness of " ^ quote (string_of_cterm cset));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    42
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    43
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1264
diff changeset
    44
(* ext_typedef *)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    45
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1264
diff changeset
    46
fun ext_typedef prep_term name (t, vs, mx) raw_set axms thms usr_tac thy =
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    47
  let
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1264
diff changeset
    48
    val dummy = require_thy thy "Set" "typedef definitions";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    49
    val sign = sign_of thy;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    50
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    51
    (*rhs*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    52
    val cset = prep_term sign raw_set;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    53
    val {T = setT, t = set, ...} = rep_cterm cset;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    54
    val rhs_tfrees = term_tfrees set;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
    val oldT = dest_setT setT handle TYPE _ =>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    56
      error ("Not a set type: " ^ quote (Sign.string_of_typ sign setT));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    57
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    58
    (*lhs*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    59
    val lhs_tfrees =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    60
      map (fn v => (v, if_none (assoc (rhs_tfrees, v)) termS)) vs;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    61
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    62
    val tname = type_name t mx;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    63
    val tlen = length vs;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    64
    val newT = Type (tname, map TFree lhs_tfrees);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    65
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    66
    val Rep_name = "Rep_" ^ name;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    67
    val Abs_name = "Abs_" ^ name;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    68
    val setC = Const (name, setT);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    69
    val RepC = Const (Rep_name, newT --> oldT);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    70
    val AbsC = Const (Abs_name, oldT --> newT);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    71
    val x_new = Free ("x", newT);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    72
    val y_old = Free ("y", oldT);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    73
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    74
    (*axioms*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    75
    val rep_type = mk_Trueprop (mk_mem (RepC $ x_new, setC));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    76
    val rep_type_inv = mk_Trueprop (mk_eq (AbsC $ (RepC $ x_new), x_new));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    77
    val abs_type_inv = mk_implies (mk_Trueprop (mk_mem (y_old, setC)),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    78
      mk_Trueprop (mk_eq (RepC $ (AbsC $ y_old), y_old)));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    79
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    80
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    81
    (* errors *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    82
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    83
    val show_names = commas_quote o map fst;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    84
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    85
    val illegal_vars =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    86
      if null (term_vars set) andalso null (term_tvars set) then []
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    87
      else ["Illegal schematic variable(s) on rhs"];
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    88
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    89
    val dup_lhs_tfrees =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    90
      (case duplicates lhs_tfrees of [] => []
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    91
      | dups => ["Duplicate type variables on lhs: " ^ show_names dups]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    92
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    93
    val extra_rhs_tfrees =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    94
      (case gen_rems (op =) (rhs_tfrees, lhs_tfrees) of [] => []
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    95
      | extras => ["Extra type variables on rhs: " ^ show_names extras]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    96
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    97
    val illegal_frees =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    98
      (case term_frees set of [] => []
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    99
      | xs => ["Illegal variables on rhs: " ^ show_names (map dest_Free xs)]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   100
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   101
    val errs = illegal_vars @ dup_lhs_tfrees @ extra_rhs_tfrees @ illegal_frees;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   102
  in
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   103
    if null errs then ()
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   104
    else error (cat_lines errs);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   105
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   106
    prove_nonempty cset (map (get_axiom thy) axms @ thms) usr_tac;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   107
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   108
    thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   109
    |> add_types [(t, tlen, mx)]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   110
    |> add_arities
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   111
     [(tname, replicate tlen logicS, logicS),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   112
      (tname, replicate tlen termS, termS)]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   113
    |> add_consts_i
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   114
     [(name, setT, NoSyn),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   115
      (Rep_name, newT --> oldT, NoSyn),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   116
      (Abs_name, oldT --> newT, NoSyn)]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   117
    |> add_defs_i
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   118
     [(name ^ "_def", mk_equals (setC, set))]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   119
    |> add_axioms_i
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   120
     [(Rep_name, rep_type),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   121
      (Rep_name ^ "_inverse", rep_type_inv),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   122
      (Abs_name ^ "_inverse", abs_type_inv)]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   123
  end
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   124
  handle ERROR =>
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1264
diff changeset
   125
    error ("The error(s) above occurred in typedef definition " ^ quote name);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   126
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   127
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   128
(* external interfaces *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   129
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   130
fun cert_term sg tm =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   131
  cterm_of sg tm handle TERM (msg, _) => error msg;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   132
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   133
fun read_term sg str =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   134
  read_cterm sg (str, termTVar);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   135
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1264
diff changeset
   136
val add_typedef = ext_typedef read_term;
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1264
diff changeset
   137
val add_typedef_i = ext_typedef cert_term;
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   138
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   139
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   140
end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   141