| author | wenzelm | 
| Thu, 01 Sep 2016 13:42:53 +0200 | |
| changeset 63744 | a406d7ab54ce | 
| parent 61268 | abe08fb15a12 | 
| child 64556 | 851ae0e7b09c | 
| permissions | -rw-r--r-- | 
| 12189 | 1 | (* Title: ZF/Tools/typechk.ML | 
| 6049 | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 6153 | 3 | Copyright 1999 University of Cambridge | 
| 6049 | 4 | |
| 18736 | 5 | Automated type checking (cf. CTT). | 
| 6049 | 6 | *) | 
| 7 | ||
| 12189 | 8 | signature TYPE_CHECK = | 
| 9 | sig | |
| 21506 | 10 | val print_tcset: Proof.context -> unit | 
| 18736 | 11 | val TC_add: attribute | 
| 12 | val TC_del: attribute | |
| 13 | val typecheck_tac: Proof.context -> tactic | |
| 14 | val type_solver_tac: Proof.context -> thm list -> int -> tactic | |
| 15 | val type_solver: solver | |
| 12189 | 16 | end; | 
| 17 | ||
| 18 | structure TypeCheck: TYPE_CHECK = | |
| 6153 | 19 | struct | 
| 12189 | 20 | |
| 18736 | 21 | (* datatype tcset *) | 
| 22 | ||
| 23 | datatype tcset = TC of | |
| 24 |  {rules: thm list,     (*the type-checking rules*)
 | |
| 25 | net: thm Net.net}; (*discrimination net of the same rules*) | |
| 26 | ||
| 42439 
9efdd0af15ac
eliminated Display.string_of_thm_without_context;
 wenzelm parents: 
38522diff
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: 
21506diff
changeset | 28 | if member Thm.eq_thm_prop rules th then | 
| 61268 | 29 |     (warning ("Ignoring duplicate type-checking rule\n" ^ Thm.string_of_thm ctxt th); tcs)
 | 
| 18736 | 30 | else | 
| 42439 
9efdd0af15ac
eliminated Display.string_of_thm_without_context;
 wenzelm parents: 
38522diff
changeset | 31 |     TC {rules = th :: rules, net = Net.insert_term (K false) (Thm.concl_of th, th) net};
 | 
| 18736 | 32 | |
| 42439 
9efdd0af15ac
eliminated Display.string_of_thm_without_context;
 wenzelm parents: 
38522diff
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: 
21506diff
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: 
21506diff
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: 
38522diff
changeset | 36 | rules = remove Thm.eq_thm_prop th rules} | 
| 61268 | 37 |   else (warning ("No such type-checking rule\n" ^ Thm.string_of_thm ctxt th); tcs);
 | 
| 6153 | 38 | |
| 39 | ||
| 18736 | 40 | (* generic data *) | 
| 41 | ||
| 33519 | 42 | structure Data = Generic_Data | 
| 18736 | 43 | ( | 
| 33519 | 44 | type T = tcset; | 
| 18736 | 45 |   val empty = TC {rules = [], net = Net.empty};
 | 
| 46 | val extend = I; | |
| 33519 | 47 |   fun merge (TC {rules, net}, TC {rules = rules', net = net'}) =
 | 
| 48 |     TC {rules = Thm.merge_thms (rules, rules'), net = Net.merge Thm.eq_thm_prop (net, net')};
 | |
| 18736 | 49 | ); | 
| 50 | ||
| 42439 
9efdd0af15ac
eliminated Display.string_of_thm_without_context;
 wenzelm parents: 
38522diff
changeset | 51 | val TC_add = | 
| 
9efdd0af15ac
eliminated Display.string_of_thm_without_context;
 wenzelm parents: 
38522diff
changeset | 52 | Thm.declaration_attribute (fn thm => fn context => | 
| 
9efdd0af15ac
eliminated Display.string_of_thm_without_context;
 wenzelm parents: 
38522diff
changeset | 53 | Data.map (add_rule (Context.proof_of context) thm) context); | 
| 
9efdd0af15ac
eliminated Display.string_of_thm_without_context;
 wenzelm parents: 
38522diff
changeset | 54 | |
| 
9efdd0af15ac
eliminated Display.string_of_thm_without_context;
 wenzelm parents: 
38522diff
changeset | 55 | val TC_del = | 
| 
9efdd0af15ac
eliminated Display.string_of_thm_without_context;
 wenzelm parents: 
38522diff
changeset | 56 | Thm.declaration_attribute (fn thm => fn context => | 
| 
9efdd0af15ac
eliminated Display.string_of_thm_without_context;
 wenzelm parents: 
38522diff
changeset | 57 | Data.map (del_rule (Context.proof_of context) thm) context); | 
| 18736 | 58 | |
| 59 | val tcset_of = Data.get o Context.Proof; | |
| 6153 | 60 | |
| 22846 | 61 | fun print_tcset ctxt = | 
| 62 |   let val TC {rules, ...} = tcset_of ctxt in
 | |
| 63 | Pretty.writeln (Pretty.big_list "type-checking rules:" | |
| 61268 | 64 | (map (Thm.pretty_thm_item ctxt) rules)) | 
| 22846 | 65 | end; | 
| 66 | ||
| 6153 | 67 | |
| 18736 | 68 | (* tactics *) | 
| 6153 | 69 | |
| 70 | (*resolution using a net rather than rules*) | |
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59164diff
changeset | 71 | fun net_res_tac ctxt maxr net = | 
| 6153 | 72 | SUBGOAL | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59164diff
changeset | 73 | (fn (prem, i) => | 
| 6153 | 74 | let val rls = Net.unify_term net (Logic.strip_assums_concl prem) | 
| 12189 | 75 | in | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59164diff
changeset | 76 | if length rls <= maxr then resolve_tac ctxt rls i else no_tac | 
| 6153 | 77 | end); | 
| 78 | ||
| 38522 | 79 | fun is_rigid_elem (Const(@{const_name Trueprop},_) $ (Const(@{const_name mem},_) $ a $ _)) =
 | 
| 6049 | 80 | not (is_Var (head_of a)) | 
| 81 | | is_rigid_elem _ = false; | |
| 82 | ||
| 12189 | 83 | (*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: 
58893diff
changeset | 84 | fun test_assume_tac ctxt = SUBGOAL(fn (prem,i) => | 
| 6049 | 85 | 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: 
58893diff
changeset | 86 | then assume_tac ctxt i else eq_assume_tac i); | 
| 6049 | 87 | |
| 12189 | 88 | (*Type checking solves a:?A (a rigid, ?A maybe flexible). | 
| 6049 | 89 | 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: 
58893diff
changeset | 90 | fun typecheck_step_tac ctxt = | 
| 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 wenzelm parents: 
58893diff
changeset | 91 |   let val TC {net, ...} = tcset_of ctxt
 | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59164diff
changeset | 92 | in FIRSTGOAL (test_assume_tac ctxt ORELSE' net_res_tac ctxt 3 net) end; | 
| 6049 | 93 | |
| 58963 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 wenzelm parents: 
58893diff
changeset | 94 | fun typecheck_tac ctxt = REPEAT (typecheck_step_tac ctxt); | 
| 6049 | 95 | |
| 59164 | 96 | (*Compile a term-net for speed*) | 
| 97 | val basic_net = | |
| 98 |   Tactic.build_net @{thms TrueI refl reflexive iff_refl ballI allI conjI impI};
 | |
| 6049 | 99 | |
| 100 | (*Instantiates variables in typing conditions. | |
| 101 | drawback: does not simplify conjunctions*) | |
| 18736 | 102 | fun type_solver_tac ctxt hyps = SELECT_GOAL | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59164diff
changeset | 103 |     (DEPTH_SOLVE (eresolve_tac ctxt @{thms FalseE} 1
 | 
| 59164 | 104 | ORELSE resolve_from_net_tac ctxt basic_net 1 | 
| 60774 | 105 | ORELSE (ares_tac ctxt hyps 1 | 
| 58963 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 wenzelm parents: 
58893diff
changeset | 106 | APPEND typecheck_step_tac ctxt))); | 
| 12189 | 107 | |
| 43596 | 108 | val type_solver = | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51658diff
changeset | 109 | Simplifier.mk_solver "ZF typecheck" (fn ctxt => | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51658diff
changeset | 110 | type_solver_tac ctxt (Simplifier.prems_of ctxt)); | 
| 6153 | 111 | |
| 58828 | 112 | val _ = Theory.setup (map_theory_simpset (fn ctxt => ctxt setSolver type_solver)); | 
| 113 | ||
| 6153 | 114 | |
| 18736 | 115 | (* concrete syntax *) | 
| 12189 | 116 | |
| 58828 | 117 | val _ = | 
| 118 | Theory.setup | |
| 119 |    (Attrib.setup @{binding TC} (Attrib.add_del TC_add TC_del)
 | |
| 120 | "declaration of type-checking rule" #> | |
| 121 |     Method.setup @{binding typecheck}
 | |
| 122 | (Method.sections | |
| 123 |         [Args.add -- Args.colon >> K (Method.modifier TC_add @{here}),
 | |
| 124 |          Args.del -- Args.colon >> K (Method.modifier TC_del @{here})]
 | |
| 125 | >> K (fn ctxt => SIMPLE_METHOD (CHANGED (typecheck_tac ctxt)))) | |
| 126 | "ZF type-checking"); | |
| 12189 | 127 | |
| 24867 | 128 | val _ = | 
| 59936 
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
 wenzelm parents: 
59498diff
changeset | 129 |   Outer_Syntax.command @{command_keyword print_tcset} "print context of ZF typecheck"
 | 
| 60097 
d20ca79d50e4
discontinued pointless warnings: commands are only defined inside a theory context;
 wenzelm parents: 
59936diff
changeset | 130 | (Scan.succeed (Toplevel.keep (print_tcset o Toplevel.context_of))); | 
| 12189 | 131 | |
| 132 | end; |