src/ZF/Tools/typechk.ML
author wenzelm
Sat, 11 Sep 2021 22:02:12 +0200
changeset 74294 ee04dc00bf0a
parent 69593 3dda49e08b9d
child 74375 ba880f3a4e52
permissions -rw-r--r--
more antiquotations;
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
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
     3
    Copyright   1999  University of Cambridge
6049
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
     4
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
     5
Automated type checking (cf. CTT).
6049
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
     6
*)
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
     7
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
     8
signature TYPE_CHECK =
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
     9
sig
21506
b2a673894ce5 prefer Proof.context over Context.generic;
wenzelm
parents: 20193
diff changeset
    10
  val print_tcset: Proof.context -> unit
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    11
  val TC_add: attribute
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    12
  val TC_del: attribute
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    13
  val typecheck_tac: Proof.context -> tactic
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    14
  val type_solver_tac: Proof.context -> thm list -> int -> tactic
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    15
  val type_solver: solver
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
    16
end;
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
    17
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
    18
structure TypeCheck: TYPE_CHECK =
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    19
struct
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
    20
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    21
(* datatype tcset *)
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    22
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    23
datatype tcset = TC of
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    24
 {rules: thm list,     (*the type-checking rules*)
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    25
  net: thm Net.net};   (*discrimination net of the same rules*)
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    26
42439
9efdd0af15ac eliminated Display.string_of_thm_without_context;
wenzelm
parents: 38522
diff changeset
    27
fun add_rule ctxt 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
    28
  if member Thm.eq_thm_prop rules th then
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 60774
diff changeset
    29
    (warning ("Ignoring duplicate type-checking rule\n" ^ Thm.string_of_thm ctxt th); tcs)
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    30
  else
42439
9efdd0af15ac eliminated Display.string_of_thm_without_context;
wenzelm
parents: 38522
diff changeset
    31
    TC {rules = th :: rules, net = Net.insert_term (K false) (Thm.concl_of th, th) net};
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    32
42439
9efdd0af15ac eliminated Display.string_of_thm_without_context;
wenzelm
parents: 38522
diff changeset
    33
fun del_rule ctxt 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
    34
  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
    35
    TC {net = Net.delete_term Thm.eq_thm_prop (Thm.concl_of th, th) net,
42439
9efdd0af15ac eliminated Display.string_of_thm_without_context;
wenzelm
parents: 38522
diff changeset
    36
      rules = remove Thm.eq_thm_prop th rules}
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 60774
diff changeset
    37
  else (warning ("No such type-checking rule\n" ^ Thm.string_of_thm ctxt th); tcs);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    38
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    39
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    40
(* generic data *)
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    41
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 32170
diff changeset
    42
structure Data = Generic_Data
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    43
(
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 32170
diff changeset
    44
  type T = tcset;
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    45
  val empty = TC {rules = [], net = Net.empty};
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    46
  val extend = I;
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 32170
diff changeset
    47
  fun merge (TC {rules, net}, TC {rules = rules', net = net'}) =
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 32170
diff changeset
    48
    TC {rules = Thm.merge_thms (rules, rules'), net = Net.merge Thm.eq_thm_prop (net, net')};
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    49
);
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    50
42439
9efdd0af15ac eliminated Display.string_of_thm_without_context;
wenzelm
parents: 38522
diff changeset
    51
val TC_add =
9efdd0af15ac eliminated Display.string_of_thm_without_context;
wenzelm
parents: 38522
diff changeset
    52
  Thm.declaration_attribute (fn thm => fn context =>
67645
5e0c81750441 trim context of persistent data;
wenzelm
parents: 64556
diff changeset
    53
    Data.map (add_rule (Context.proof_of context) (Thm.trim_context thm)) context);
42439
9efdd0af15ac eliminated Display.string_of_thm_without_context;
wenzelm
parents: 38522
diff changeset
    54
9efdd0af15ac eliminated Display.string_of_thm_without_context;
wenzelm
parents: 38522
diff changeset
    55
val TC_del =
9efdd0af15ac eliminated Display.string_of_thm_without_context;
wenzelm
parents: 38522
diff changeset
    56
  Thm.declaration_attribute (fn thm => fn context =>
9efdd0af15ac eliminated Display.string_of_thm_without_context;
wenzelm
parents: 38522
diff changeset
    57
    Data.map (del_rule (Context.proof_of context) thm) context);
18736
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:"
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 60774
diff changeset
    64
      (map (Thm.pretty_thm_item ctxt) rules))
22846
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*)
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59164
diff changeset
    71
fun net_res_tac ctxt maxr net =
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    72
  SUBGOAL
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59164
diff changeset
    73
    (fn (prem, i) =>
6153
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
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59164
diff changeset
    76
         if length rls <= maxr then resolve_tac ctxt 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
74294
ee04dc00bf0a more antiquotations;
wenzelm
parents: 69593
diff changeset
    79
fun is_rigid_elem \<^Const_>\<open>Trueprop for \<open>\<^Const_>\<open>mem for a _\<close>\<close>\<close> = not (is_Var (head_of a))
6049
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
    80
  | is_rigid_elem _ = false;
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
    81
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
    82
(*Try solving a:A by assumption provided a is rigid!*)
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58893
diff changeset
    83
fun test_assume_tac ctxt = SUBGOAL(fn (prem,i) =>
6049
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
    84
    if is_rigid_elem (Logic.strip_assums_concl prem)
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58893
diff changeset
    85
    then  assume_tac ctxt i  else  eq_assume_tac i);
6049
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
    86
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
    87
(*Type checking solves a:?A (a rigid, ?A maybe flexible).
6049
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
    88
  match_tac is too strict; would refuse to instantiate ?A*)
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58893
diff changeset
    89
fun typecheck_step_tac ctxt =
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58893
diff changeset
    90
  let val TC {net, ...} = tcset_of ctxt
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59164
diff changeset
    91
  in FIRSTGOAL (test_assume_tac ctxt ORELSE' net_res_tac ctxt 3 net) end;
6049
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
    92
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58893
diff changeset
    93
fun typecheck_tac ctxt = REPEAT (typecheck_step_tac ctxt);
6049
7fef0169ab5e moved from ZF to new subdirectory Tools
paulson
parents:
diff changeset
    94
59164
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 58963
diff changeset
    95
(*Compile a term-net for speed*)
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 58963
diff changeset
    96
val basic_net =
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 58963
diff changeset
    97
  Tactic.build_net @{thms TrueI refl reflexive iff_refl ballI allI conjI 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
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59164
diff changeset
   102
    (DEPTH_SOLVE (eresolve_tac ctxt @{thms FalseE} 1
59164
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 58963
diff changeset
   103
                  ORELSE resolve_from_net_tac ctxt basic_net 1
60774
6c28d8ed2488 proper context;
wenzelm
parents: 60097
diff changeset
   104
                  ORELSE (ares_tac ctxt hyps 1
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58893
diff changeset
   105
                          APPEND typecheck_step_tac ctxt)));
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
   106
43596
78211f66cf8d simplified/unified Simplifier.mk_solver;
wenzelm
parents: 42795
diff changeset
   107
val type_solver =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51658
diff changeset
   108
  Simplifier.mk_solver "ZF typecheck" (fn ctxt =>
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51658
diff changeset
   109
    type_solver_tac ctxt (Simplifier.prems_of ctxt));
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   110
58828
6d076fdd933d modernized setup;
wenzelm
parents: 58048
diff changeset
   111
val _ = Theory.setup (map_theory_simpset (fn ctxt => ctxt setSolver type_solver));
6d076fdd933d modernized setup;
wenzelm
parents: 58048
diff changeset
   112
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   113
18736
541d04c79e12 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   114
(* concrete syntax *)
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
   115
58828
6d076fdd933d modernized setup;
wenzelm
parents: 58048
diff changeset
   116
val _ =
6d076fdd933d modernized setup;
wenzelm
parents: 58048
diff changeset
   117
  Theory.setup
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67645
diff changeset
   118
   (Attrib.setup \<^binding>\<open>TC\<close> (Attrib.add_del TC_add TC_del)
58828
6d076fdd933d modernized setup;
wenzelm
parents: 58048
diff changeset
   119
      "declaration of type-checking rule" #>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67645
diff changeset
   120
    Method.setup \<^binding>\<open>typecheck\<close>
58828
6d076fdd933d modernized setup;
wenzelm
parents: 58048
diff changeset
   121
      (Method.sections
64556
851ae0e7b09c more symbols;
wenzelm
parents: 61268
diff changeset
   122
        [Args.add -- Args.colon >> K (Method.modifier TC_add \<^here>),
851ae0e7b09c more symbols;
wenzelm
parents: 61268
diff changeset
   123
         Args.del -- Args.colon >> K (Method.modifier TC_del \<^here>)]
58828
6d076fdd933d modernized setup;
wenzelm
parents: 58048
diff changeset
   124
        >> K (fn ctxt => SIMPLE_METHOD (CHANGED (typecheck_tac ctxt))))
6d076fdd933d modernized setup;
wenzelm
parents: 58048
diff changeset
   125
      "ZF type-checking");
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
   126
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24826
diff changeset
   127
val _ =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67645
diff changeset
   128
  Outer_Syntax.command \<^command_keyword>\<open>print_tcset\<close> "print context of ZF typecheck"
60097
d20ca79d50e4 discontinued pointless warnings: commands are only defined inside a theory context;
wenzelm
parents: 59936
diff changeset
   129
    (Scan.succeed (Toplevel.keep (print_tcset o Toplevel.context_of)));
12189
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
   130
4729bbf86626 Isar attribute and method setup;
wenzelm
parents: 12109
diff changeset
   131
end;