| author | paulson | 
| Tue, 22 Sep 1998 15:24:39 +0200 | |
| changeset 5533 | bce36a019b03 | 
| parent 4352 | 7ac9f3e8a97d | 
| permissions | -rw-r--r-- | 
| 1461 | 1 | (* Title: ZF/constructor.ML | 
| 0 | 2 | ID: $Id$ | 
| 1461 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 0 | 4 | Copyright 1993 University of Cambridge | 
| 5 | ||
| 516 | 6 | Constructor function module -- for (Co)Datatype Definitions | 
| 0 | 7 | *) | 
| 8 | ||
| 516 | 9 | signature CONSTRUCTOR_ARG = | 
| 0 | 10 | sig | 
| 1461 | 11 | val thy : theory (*parent containing constructor defs*) | 
| 12 | val big_rec_name : string (*name of mutually recursive set*) | |
| 516 | 13 | val con_ty_lists : ((string*typ*mixfix) * | 
| 1461 | 14 | string * term list * term list) list list | 
| 15 | (*description of constructors*) | |
| 0 | 16 | end; | 
| 17 | ||
| 18 | signature CONSTRUCTOR_RESULT = | |
| 19 | sig | |
| 1461 | 20 | val con_defs : thm list (*definitions made in thy*) | 
| 21 | val case_eqns : thm list (*equations for case operator*) | |
| 22 | val free_iffs : thm list (*freeness rewrite rules*) | |
| 23 | val free_SEs : thm list (*freeness destruct rules*) | |
| 24 | val mk_free : string -> thm (*makes freeness theorems*) | |
| 0 | 25 | end; | 
| 26 | ||
| 27 | ||
| 516 | 28 | (*Proves theorems relating to constructors*) | 
| 29 | functor Constructor_Fun (structure Const: CONSTRUCTOR_ARG and | |
| 0 | 30 | Pr : PR and Su : SU) : CONSTRUCTOR_RESULT = | 
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1103diff
changeset | 31 | let | 
| 0 | 32 | |
| 516 | 33 | (*1st element is the case definition; others are the constructors*) | 
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1103diff
changeset | 34 | val big_case_name = Const.big_rec_name ^ "_case"; | 
| 0 | 35 | |
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1103diff
changeset | 36 | val con_defs = get_def Const.thy big_case_name :: | 
| 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1103diff
changeset | 37 | map (get_def Const.thy o #2) (flat Const.con_ty_lists); | 
| 0 | 38 | |
| 39 | (** Prove the case theorem **) | |
| 40 | ||
| 516 | 41 | (*Get the case term from its definition*) | 
| 42 | val Const("==",_) $ big_case_tm $ _ =
 | |
| 543 
e961b2092869
ZF/ind_syntax/unvarifyT, unvarify: moved to Pure/logic.ML
 lcp parents: 
516diff
changeset | 43 | hd con_defs |> rep_thm |> #prop |> Logic.unvarify; | 
| 516 | 44 | |
| 45 | val (_, big_case_args) = strip_comb big_case_tm; | |
| 46 | ||
| 0 | 47 | (*Each equation has the form | 
| 48 | rec_case(f_con1,...,f_conn)(coni(args)) = f_coni(args) *) | |
| 516 | 49 | fun mk_case_equation (((id,T,syn), name, args, prems), case_free) = | 
| 4352 
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
 paulson parents: 
3925diff
changeset | 50 | FOLogic.mk_Trueprop | 
| 
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
 paulson parents: 
3925diff
changeset | 51 | (FOLogic.mk_eq | 
| 
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
 paulson parents: 
3925diff
changeset | 52 | (big_case_tm $ | 
| 
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
 paulson parents: 
3925diff
changeset | 53 | (list_comb (Const (Sign.intern_const (sign_of Const.thy) name,T), | 
| 
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
 paulson parents: 
3925diff
changeset | 54 | args)), | 
| 
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
 paulson parents: 
3925diff
changeset | 55 | list_comb (case_free, args))); | 
| 0 | 56 | |
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1103diff
changeset | 57 | val case_trans = hd con_defs RS Ind_Syntax.def_trans | 
| 1103 
08fda5148971
prove_case_equation now calls uses meta_eq_to_obj_eq to cope
 lcp parents: 
543diff
changeset | 58 | and split_trans = Pr.split_eq RS meta_eq_to_obj_eq RS trans; | 
| 0 | 59 | |
| 516 | 60 | (*Proves a single case equation. Could use simp_tac, but it's slower!*) | 
| 0 | 61 | fun case_tacsf con_def _ = | 
| 62 | [rewtac con_def, | |
| 63 | rtac case_trans 1, | |
| 1103 
08fda5148971
prove_case_equation now calls uses meta_eq_to_obj_eq to cope
 lcp parents: 
543diff
changeset | 64 | REPEAT (resolve_tac [refl, split_trans, | 
| 
08fda5148971
prove_case_equation now calls uses meta_eq_to_obj_eq to cope
 lcp parents: 
543diff
changeset | 65 | Su.case_inl RS trans, | 
| 
08fda5148971
prove_case_equation now calls uses meta_eq_to_obj_eq to cope
 lcp parents: 
543diff
changeset | 66 | Su.case_inr RS trans] 1)]; | 
| 0 | 67 | |
| 68 | fun prove_case_equation (arg,con_def) = | |
| 543 
e961b2092869
ZF/ind_syntax/unvarifyT, unvarify: moved to Pure/logic.ML
 lcp parents: 
516diff
changeset | 69 | prove_goalw_cterm [] | 
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1103diff
changeset | 70 | (cterm_of (sign_of Const.thy) (mk_case_equation arg)) | 
| 1103 
08fda5148971
prove_case_equation now calls uses meta_eq_to_obj_eq to cope
 lcp parents: 
543diff
changeset | 71 | (case_tacsf con_def); | 
| 0 | 72 | |
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1103diff
changeset | 73 | val con_iffs = con_defs RL [Ind_Syntax.def_swap_iff]; | 
| 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1103diff
changeset | 74 | |
| 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1103diff
changeset | 75 | in | 
| 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1103diff
changeset | 76 | struct | 
| 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1103diff
changeset | 77 | val con_defs = con_defs | 
| 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1103diff
changeset | 78 | |
| 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1103diff
changeset | 79 | val free_iffs = con_iffs @ | 
| 0 | 80 | [Su.distinct, Su.distinct', Su.inl_iff, Su.inr_iff, Pr.pair_iff]; | 
| 81 | ||
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1103diff
changeset | 82 | val free_SEs = Ind_Syntax.mk_free_SEs con_iffs @ Su.free_SEs; | 
| 0 | 83 | |
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1103diff
changeset | 84 | (*Typical theorems have the form ~con1=con2, con1=con2==>False, | 
| 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1103diff
changeset | 85 | con1(x)=con1(y) ==> x=y, con1(x)=con1(y) <-> x=y, etc. *) | 
| 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1103diff
changeset | 86 | fun mk_free s = | 
| 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1103diff
changeset | 87 | prove_goalw Const.thy con_defs s | 
| 1461 | 88 | (fn prems => [cut_facts_tac prems 1, | 
| 89 | fast_tac (ZF_cs addSEs free_SEs) 1]); | |
| 0 | 90 | |
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1103diff
changeset | 91 | val case_eqns = map prove_case_equation | 
| 1461 | 92 | (flat Const.con_ty_lists ~~ big_case_args ~~ tl con_defs); | 
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1103diff
changeset | 93 | end | 
| 0 | 94 | end; | 
| 95 | ||
| 96 |