src/ZF/Tools/typechk.ML
author wenzelm
Fri, 13 Mar 2009 19:58:26 +0100
changeset 30510 4120fc59dd85
parent 26496 49ae9456eba9
child 30528 7173bf123335
permissions -rw-r--r--
unified type Proof.method and pervasive METHOD combinators;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
     1
(*  Title:      ZF/Tools/typechk.ML
6049
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
     2
    ID:         $Id$
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
     4
    Copyright   1999  University of Cambridge
6049
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
     5
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
     6
Automated type checking (cf. CTT).
6049
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
     7
*)
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
     8
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
     9
signature TYPE_CHECK =
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
    10
sig
21506
b2a673894ce5 prefer Proof.context over Context.generic;
wenzelm
parents: 20193
diff changeset
    11
  val print_tcset: Proof.context -> unit
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    12
  val TC_add: attribute
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    13
  val TC_del: attribute
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    14
  val typecheck_tac: Proof.context -> tactic
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    15
  val type_solver_tac: Proof.context -> thm list -> int -> tactic
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    16
  val type_solver: solver
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 17886
diff changeset
    17
  val setup: theory -> theory
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
    18
end;
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
    19
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
    20
structure TypeCheck: TYPE_CHECK =
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    21
struct
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
    22
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    23
(* datatype tcset *)
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    24
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    25
datatype tcset = TC of
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    26
 {rules: thm list,     (*the type-checking rules*)
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    27
  net: thm Net.net};   (*discrimination net of the same rules*)
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    28
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    29
fun add_rule th (tcs as TC {rules, net}) =
22360
26ead7ed4f4b moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents: 21506
diff changeset
    30
  if member Thm.eq_thm_prop rules th then
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    31
    (warning ("Ignoring duplicate type-checking rule\n" ^ Display.string_of_thm th); tcs)
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    32
  else
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    33
    TC {rules = th :: rules,
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    34
        net = Net.insert_term (K false) (Thm.concl_of th, th) net};
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    35
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    36
fun del_rule th (tcs as TC {rules, net}) =
22360
26ead7ed4f4b moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents: 21506
diff changeset
    37
  if member Thm.eq_thm_prop rules th then
26ead7ed4f4b moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents: 21506
diff changeset
    38
    TC {net = Net.delete_term Thm.eq_thm_prop (Thm.concl_of th, th) net,
26ead7ed4f4b moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents: 21506
diff changeset
    39
        rules = remove Thm.eq_thm_prop th rules}
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    40
  else (warning ("No such type-checking rule\n" ^ Display.string_of_thm th); tcs);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    41
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    42
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    43
(* generic data *)
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    44
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    45
structure Data = GenericDataFun
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    46
(
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    47
  type T = tcset
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    48
  val empty = TC {rules = [], net = Net.empty};
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    49
  val extend = I;
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    50
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    51
  fun merge _ (TC {rules, net}, TC {rules = rules', net = net'}) =
24039
273698405054 renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 22846
diff changeset
    52
    TC {rules = Thm.merge_thms (rules, rules'),
22360
26ead7ed4f4b moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents: 21506
diff changeset
    53
        net = Net.merge Thm.eq_thm_prop (net, net')};
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    54
);
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    55
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    56
val TC_add = Thm.declaration_attribute (Data.map o add_rule);
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    57
val TC_del = Thm.declaration_attribute (Data.map o del_rule);
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    58
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    59
val tcset_of = Data.get o Context.Proof;
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    60
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22360
diff changeset
    61
fun print_tcset ctxt =
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22360
diff changeset
    62
  let val TC {rules, ...} = tcset_of ctxt in
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22360
diff changeset
    63
    Pretty.writeln (Pretty.big_list "type-checking rules:"
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22360
diff changeset
    64
      (map (ProofContext.pretty_thm ctxt) rules))
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22360
diff changeset
    65
  end;
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22360
diff changeset
    66
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    67
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    68
(* tactics *)
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    69
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    70
(*resolution using a net rather than rules*)
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    71
fun net_res_tac maxr net =
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    72
  SUBGOAL
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    73
    (fn (prem,i) =>
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    74
      let val rls = Net.unify_term net (Logic.strip_assums_concl prem)
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
    75
      in
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
    76
         if length rls <= maxr then resolve_tac rls i else no_tac
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    77
      end);
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    78
24826
78e6a3cea367 avoid unnamed infixes;
wenzelm
parents: 24039
diff changeset
    79
fun is_rigid_elem (Const("Trueprop",_) $ (Const(@{const_name mem},_) $ a $ _)) =
6049
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
    80
      not (is_Var (head_of a))
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
    81
  | is_rigid_elem _ = false;
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
    82
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
    83
(*Try solving a:A by assumption provided a is rigid!*)
6049
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
    84
val test_assume_tac = SUBGOAL(fn (prem,i) =>
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
    85
    if is_rigid_elem (Logic.strip_assums_concl prem)
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
    86
    then  assume_tac i  else  eq_assume_tac i);
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
    87
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
    88
(*Type checking solves a:?A (a rigid, ?A maybe flexible).
6049
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
    89
  match_tac is too strict; would refuse to instantiate ?A*)
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    90
fun typecheck_step_tac (TC{net,...}) =
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    91
    FIRSTGOAL (test_assume_tac ORELSE' net_res_tac 3 net);
6049
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
    92
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    93
fun typecheck_tac ctxt = REPEAT (typecheck_step_tac (tcset_of ctxt));
6049
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
    94
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    95
(*Compiles a term-net for speed*)
26287
df8e5362cff9 proper antiquotations;
wenzelm
parents: 24893
diff changeset
    96
val basic_res_tac = net_resolve_tac [@{thm TrueI}, @{thm refl}, reflexive_thm, @{thm iff_refl},
df8e5362cff9 proper antiquotations;
wenzelm
parents: 24893
diff changeset
    97
                                     @{thm ballI}, @{thm allI}, @{thm conjI}, @{thm impI}];
6049
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
    98
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
    99
(*Instantiates variables in typing conditions.
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
   100
  drawback: does not simplify conjunctions*)
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   101
fun type_solver_tac ctxt hyps = SELECT_GOAL
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   102
    (DEPTH_SOLVE (etac FalseE 1
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
   103
                  ORELSE basic_res_tac 1
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
   104
                  ORELSE (ares_tac hyps 1
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   105
                          APPEND typecheck_step_tac (tcset_of ctxt))));
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
   106
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   107
val type_solver = Simplifier.mk_solver' "ZF typecheck" (fn ss =>
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   108
  type_solver_tac (Simplifier.the_context ss) (Simplifier.prems_of_ss ss));
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   109
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   110
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   111
(* concrete syntax *)
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
   112
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   113
val TC_att = Attrib.add_del_args TC_add TC_del;
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
   114
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   115
val typecheck_meth =
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   116
  Method.only_sectioned_args
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   117
    [Args.add -- Args.colon >> K (I, TC_add),
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   118
     Args.del -- Args.colon >> K (I, TC_del)]
30510
4120fc59dd85 unified type Proof.method and pervasive METHOD combinators;
wenzelm
parents: 26496
diff changeset
   119
  (fn ctxt => SIMPLE_METHOD (CHANGED (typecheck_tac ctxt)));
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
   120
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24826
diff changeset
   121
val _ =
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24826
diff changeset
   122
  OuterSyntax.improper_command "print_tcset" "print context of ZF typecheck" OuterKeyword.diag
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   123
    (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_context o
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24826
diff changeset
   124
      Toplevel.keep (print_tcset o Toplevel.context_of)));
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
   125
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
   126
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   127
(* theory setup *)
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
   128
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
   129
val setup =
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   130
  Attrib.add_attributes [("TC", TC_att, "declaration of type-checking rule")] #>
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   131
  Method.add_methods [("typecheck", typecheck_meth, "ZF type-checking")] #>
26496
49ae9456eba9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26287
diff changeset
   132
  Simplifier.map_simpset (fn ss => ss setSolver type_solver);
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
   133
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
   134
end;