| author | lcp | 
| Fri, 16 Dec 1994 17:41:49 +0100 | |
| changeset 800 | 23f55b829ccb | 
| parent 762 | 1cf9ebcc3ff3 | 
| child 855 | 4c8d0ece1f95 | 
| permissions | -rw-r--r-- | 
| 0 | 1 | (* Title: ZF/simpdata | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1991 University of Cambridge | |
| 5 | ||
| 6 | Rewriting for ZF set theory -- based on FOL rewriting | |
| 7 | *) | |
| 8 | ||
| 9 | (** Tactics for type checking -- from CTT **) | |
| 10 | ||
| 11 | fun is_rigid_elem (Const("Trueprop",_) $ (Const("op :",_) $ a $ _)) = 
 | |
| 12 | not (is_Var (head_of a)) | |
| 13 | | is_rigid_elem _ = false; | |
| 14 | ||
| 15 | (*Try solving a:A by assumption provided a is rigid!*) | |
| 16 | val test_assume_tac = SUBGOAL(fn (prem,i) => | |
| 17 | if is_rigid_elem (Logic.strip_assums_concl prem) | |
| 762 
1cf9ebcc3ff3
test_assume_tac: now tries eq_assume_tac on exceptional cases
 lcp parents: 
533diff
changeset | 18 | then assume_tac i else eq_assume_tac i); | 
| 0 | 19 | |
| 20 | (*Type checking solves a:?A (a rigid, ?A maybe flexible). | |
| 21 | match_tac is too strict; would refuse to instantiate ?A*) | |
| 22 | fun typechk_step_tac tyrls = | |
| 23 | FIRSTGOAL (test_assume_tac ORELSE' filt_resolve_tac tyrls 3); | |
| 24 | ||
| 25 | fun typechk_tac tyrls = REPEAT (typechk_step_tac tyrls); | |
| 26 | ||
| 27 | val ZF_typechecks = [if_type,lam_type,SigmaI,apply_type,split_type]; | |
| 28 | ||
| 516 | 29 | (*Instantiates variables in typing conditions. | 
| 30 | drawback: does not simplify conjunctions*) | |
| 0 | 31 | fun type_auto_tac tyrls hyps = SELECT_GOAL | 
| 32 | (DEPTH_SOLVE (typechk_step_tac (tyrls@hyps) | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 33 | ORELSE ares_tac [TrueI,refl,iff_refl,ballI,allI,conjI,impI] 1)); | 
| 0 | 34 | |
| 35 | (** New version of mk_rew_rules **) | |
| 36 | ||
| 37 | (*Should False yield False<->True, or should it solve goals some other way?*) | |
| 38 | ||
| 39 | (*Analyse a rigid formula*) | |
| 40 | val atomize_pairs = | |
| 41 |   [("Ball",	[bspec]), 
 | |
| 42 |    ("All",	[spec]),
 | |
| 43 |    ("op -->",	[mp]),
 | |
| 44 |    ("op &",	[conjunct1,conjunct2])];
 | |
| 45 | ||
| 46 | (*Analyse a:b, where b is rigid*) | |
| 47 | val atomize_mem_pairs = | |
| 48 |   [("Collect",	[CollectD1,CollectD2]),
 | |
| 49 |    ("op -",	[DiffD1,DiffD2]),
 | |
| 50 |    ("op Int",	[IntD1,IntD2])];
 | |
| 51 | ||
| 52 | (*Analyse a theorem to atomic rewrite rules*) | |
| 53 | fun atomize th = | |
| 54 | let fun tryrules pairs t = | |
| 55 | case head_of t of | |
| 56 | Const(a,_) => | |
| 57 | (case assoc(pairs,a) of | |
| 58 | Some rls => flat (map atomize ([th] RL rls)) | |
| 59 | | None => [th]) | |
| 60 | | _ => [th] | |
| 435 | 61 | in case concl_of th of | 
| 62 |        Const("Trueprop",_) $ P => 
 | |
| 63 | (case P of | |
| 64 | 	       Const("op :",_) $ a $ b => tryrules atomize_mem_pairs b
 | |
| 65 | 	     | Const("True",_)         => []
 | |
| 66 | 	     | Const("False",_)        => []
 | |
| 67 | | A => tryrules atomize_pairs A) | |
| 68 | | _ => [th] | |
| 0 | 69 | end; | 
| 70 | ||
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 71 | val ZF_simps = [empty_subsetI, consI1, succI1, ball_simp, if_true, if_false, | 
| 467 | 72 | beta, eta, restrict, fst_conv, snd_conv, split, Pair_iff, | 
| 485 | 73 | triv_RepFun, subset_refl]; | 
| 0 | 74 | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 75 | (*Sigma_cong, Pi_cong NOT included by default since they cause | 
| 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 76 | flex-flex pairs and the "Check your prover" error -- because most | 
| 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 77 | Sigma's and Pi's are abbreviated as * or -> *) | 
| 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 78 | val ZF_congs = | 
| 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 79 | [ball_cong, bex_cong, Replace_cong, RepFun_cong, Collect_cong, lam_cong]; | 
| 0 | 80 | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 81 | val ZF_ss = FOL_ss | 
| 279 | 82 | setmksimps (map mk_meta_eq o atomize o gen_all) | 
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 83 | addsimps (ZF_simps @ mem_simps @ bquant_simps) | 
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 84 | addcongs ZF_congs; |