src/HOL/thy_data.ML
author paulson
Thu, 08 May 1997 10:20:37 +0200
changeset 3142 a6f73a02c619
parent 3040 7d48671753da
child 3265 8358e19d0d4c
permissions -rw-r--r--
Made a slow proof slightly faster
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1455
diff changeset
     1
(*  Title:      HOL/thy_data.ML
1455
52a0271621f3 changed the way simpsets and information about datatypes are stored
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1455
diff changeset
     3
    Author:     Carsten Clasohm
1455
52a0271621f3 changed the way simpsets and information about datatypes are stored
clasohm
parents:
diff changeset
     4
    Copyright   1995 TU Muenchen
52a0271621f3 changed the way simpsets and information about datatypes are stored
clasohm
parents:
diff changeset
     5
52a0271621f3 changed the way simpsets and information about datatypes are stored
clasohm
parents:
diff changeset
     6
Definitions that have to be reread after init_thy_reader has been invoked
52a0271621f3 changed the way simpsets and information about datatypes are stored
clasohm
parents:
diff changeset
     7
*)
52a0271621f3 changed the way simpsets and information about datatypes are stored
clasohm
parents:
diff changeset
     8
52a0271621f3 changed the way simpsets and information about datatypes are stored
clasohm
parents:
diff changeset
     9
fun simpset_of tname =
1657
a7a6c3fb3cdd changed first parameter of add_thydata and get_thydata
clasohm
parents: 1465
diff changeset
    10
  case get_thydata tname "simpset" of
1455
52a0271621f3 changed the way simpsets and information about datatypes are stored
clasohm
parents:
diff changeset
    11
      None => empty_ss
52a0271621f3 changed the way simpsets and information about datatypes are stored
clasohm
parents:
diff changeset
    12
    | Some (SS_DATA ss) => ss;
52a0271621f3 changed the way simpsets and information about datatypes are stored
clasohm
parents:
diff changeset
    13
1726
cbea219f5e5a Added function claset_of.
berghofe
parents: 1668
diff changeset
    14
fun claset_of tname =
cbea219f5e5a Added function claset_of.
berghofe
parents: 1668
diff changeset
    15
  case get_thydata tname "claset" of
cbea219f5e5a Added function claset_of.
berghofe
parents: 1668
diff changeset
    16
      None => empty_cs
cbea219f5e5a Added function claset_of.
berghofe
parents: 1668
diff changeset
    17
    | Some (CS_DATA cs) => cs;
cbea219f5e5a Added function claset_of.
berghofe
parents: 1668
diff changeset
    18
1668
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    19
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    20
(** Information about datatypes **)
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    21
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    22
(* Retrieve information for a specific datatype *)
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    23
fun datatype_info thy tname =
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    24
  case get_thydata (thyname_of_sign (sign_of thy)) "datatypes" of
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    25
      None => None
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    26
    | Some (DT_DATA ds) => assoc (ds, tname);
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    27
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    28
fun match_info thy tname =
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    29
  case datatype_info thy tname of
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    30
      Some {case_const, constructors, ...} =>
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    31
        {case_const=case_const, constructors=constructors}
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    32
    | None => error ("No datatype \"" ^ tname ^ "\" defined in this theory.");
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    33
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    34
fun induct_info thy tname =
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    35
  case datatype_info thy tname of
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    36
      Some {constructors, nchotomy, ...} =>
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    37
        {constructors=constructors, nchotomy=nchotomy}
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    38
    | None => error ("No datatype \"" ^ tname ^ "\" defined in this theory.");
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    39
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    40
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    41
(* Retrieve information for all datatypes defined in a theory and its
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    42
   ancestors *)
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    43
fun extract_info thy =
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    44
  let val (congs, rewrites) =
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    45
        case get_thydata (thyname_of_sign (sign_of thy)) "datatypes" of
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    46
            None => ([], [])
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    47
          | Some (DT_DATA ds) =>
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    48
              let val info = map snd ds
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    49
              in (map #case_cong info, flat (map #case_rewrites info)) end;
8ead1fe65aad added Konrad's code for the datatype package
clasohm
parents: 1657
diff changeset
    50
  in {case_congs = congs, case_rewrites = rewrites} end;
2562
d571d6660240 Moved qed_spec_mp, etc., from HOL.ML to thy_data.ML so that they work
paulson
parents: 1726
diff changeset
    51
3040
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 2562
diff changeset
    52
(* generic induction tactic for datatypes *)
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 2562
diff changeset
    53
fun induct_tac var i =
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 2562
diff changeset
    54
  let fun find_tname Bi =
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 2562
diff changeset
    55
        let val frees = map (fn Free x => x) (term_frees Bi)
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 2562
diff changeset
    56
            val params = Logic.strip_params Bi;
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 2562
diff changeset
    57
        in case assoc (frees@params, var) of
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 2562
diff changeset
    58
             None => error("No such variable in subgoal: " ^ quote var)
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 2562
diff changeset
    59
           | Some(Type(tn,_)) => tn
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 2562
diff changeset
    60
           | _ => error("Cannot induct on type of " ^ quote var)
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 2562
diff changeset
    61
        end;
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 2562
diff changeset
    62
      fun get_ind_tac sign tn =
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 2562
diff changeset
    63
        (case get_thydata (thyname_of_sign sign) "datatypes" of
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 2562
diff changeset
    64
           None => error ("No such datatype: " ^ quote tn)
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 2562
diff changeset
    65
         | Some (DT_DATA ds) =>
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 2562
diff changeset
    66
             (case assoc (ds, tn) of
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 2562
diff changeset
    67
                Some {induct_tac, ...} => induct_tac
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 2562
diff changeset
    68
              | None => error ("Not a datatype: " ^ quote tn)));
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 2562
diff changeset
    69
      fun induct state =
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 2562
diff changeset
    70
        let val (_, _, Bi, _) = dest_state (state, i) 
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 2562
diff changeset
    71
        in get_ind_tac (#sign(rep_thm state)) (find_tname Bi) var i end
7d48671753da Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents: 2562
diff changeset
    72
  in STATE induct end;
2562
d571d6660240 Moved qed_spec_mp, etc., from HOL.ML to thy_data.ML so that they work
paulson
parents: 1726
diff changeset
    73
d571d6660240 Moved qed_spec_mp, etc., from HOL.ML to thy_data.ML so that they work
paulson
parents: 1726
diff changeset
    74
(*Must be redefined in order to refer to the new instance of bind_thm
d571d6660240 Moved qed_spec_mp, etc., from HOL.ML to thy_data.ML so that they work
paulson
parents: 1726
diff changeset
    75
  created by init_thy_reader.*)
d571d6660240 Moved qed_spec_mp, etc., from HOL.ML to thy_data.ML so that they work
paulson
parents: 1726
diff changeset
    76
d571d6660240 Moved qed_spec_mp, etc., from HOL.ML to thy_data.ML so that they work
paulson
parents: 1726
diff changeset
    77
fun qed_spec_mp name =
d571d6660240 Moved qed_spec_mp, etc., from HOL.ML to thy_data.ML so that they work
paulson
parents: 1726
diff changeset
    78
  let val thm = normalize_thm [RSspec,RSmp] (result())
d571d6660240 Moved qed_spec_mp, etc., from HOL.ML to thy_data.ML so that they work
paulson
parents: 1726
diff changeset
    79
  in bind_thm(name, thm) end;
d571d6660240 Moved qed_spec_mp, etc., from HOL.ML to thy_data.ML so that they work
paulson
parents: 1726
diff changeset
    80
d571d6660240 Moved qed_spec_mp, etc., from HOL.ML to thy_data.ML so that they work
paulson
parents: 1726
diff changeset
    81
fun qed_goal_spec_mp name thy s p = 
d571d6660240 Moved qed_spec_mp, etc., from HOL.ML to thy_data.ML so that they work
paulson
parents: 1726
diff changeset
    82
	bind_thm (name, normalize_thm [RSspec,RSmp] (prove_goal thy s p));
d571d6660240 Moved qed_spec_mp, etc., from HOL.ML to thy_data.ML so that they work
paulson
parents: 1726
diff changeset
    83
d571d6660240 Moved qed_spec_mp, etc., from HOL.ML to thy_data.ML so that they work
paulson
parents: 1726
diff changeset
    84
fun qed_goalw_spec_mp name thy defs s p = 
d571d6660240 Moved qed_spec_mp, etc., from HOL.ML to thy_data.ML so that they work
paulson
parents: 1726
diff changeset
    85
	bind_thm (name, normalize_thm [RSspec,RSmp] (prove_goalw thy defs s p));