author | wenzelm |
Fri, 04 Jan 2019 23:22:53 +0100 | |
changeset 69593 | 3dda49e08b9d |
parent 60642 | 48dd1cefb4ae |
child 74282 | c2ee8d993d6a |
permissions | -rw-r--r-- |
24584 | 1 |
(* Title: ZF/Tools/cartprod.ML |
6049 | 2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
3 |
Copyright 1996 University of Cambridge |
|
4 |
||
35762 | 5 |
Signatures for inductive definitions. |
6049 | 6 |
|
35762 | 7 |
Syntactic operations for Cartesian Products. |
6049 | 8 |
*) |
9 |
||
10 |
signature FP = (** Description of a fixed point operator **) |
|
11 |
sig |
|
12 |
val oper : term (*fixed point operator*) |
|
13 |
val bnd_mono : term (*monotonicity predicate*) |
|
14 |
val bnd_monoI : thm (*intro rule for bnd_mono*) |
|
15 |
val subs : thm (*subset theorem for fp*) |
|
16 |
val Tarski : thm (*Tarski's fixed point theorem*) |
|
17 |
val induct : thm (*induction/coinduction rule*) |
|
18 |
end; |
|
19 |
||
20 |
signature SU = (** Description of a disjoint sum **) |
|
21 |
sig |
|
22 |
val sum : term (*disjoint sum operator*) |
|
23 |
val inl : term (*left injection*) |
|
24 |
val inr : term (*right injection*) |
|
25 |
val elim : term (*case operator*) |
|
26 |
val case_inl : thm (*inl equality rule for case*) |
|
27 |
val case_inr : thm (*inr equality rule for case*) |
|
28 |
val inl_iff : thm (*injectivity of inl, using <->*) |
|
29 |
val inr_iff : thm (*injectivity of inr, using <->*) |
|
30 |
val distinct : thm (*distinctness of inl, inr using <->*) |
|
31 |
val distinct' : thm (*distinctness of inr, inl using <->*) |
|
32 |
val free_SEs : thm list (*elim rules for SU, and pair_iff!*) |
|
33 |
end; |
|
34 |
||
35 |
signature PR = (** Description of a Cartesian product **) |
|
36 |
sig |
|
37 |
val sigma : term (*Cartesian product operator*) |
|
38 |
val pair : term (*pairing operator*) |
|
39 |
val split_name : string (*name of polymorphic split*) |
|
40 |
val pair_iff : thm (*injectivity of pairing, using <->*) |
|
41 |
val split_eq : thm (*equality rule for split*) |
|
42 |
val fsplitI : thm (*intro rule for fsplit*) |
|
43 |
val fsplitD : thm (*destruct rule for fsplit*) |
|
44 |
val fsplitE : thm (*elim rule; apparently never used*) |
|
45 |
end; |
|
46 |
||
47 |
signature CARTPROD = (** Derived syntactic functions for produts **) |
|
48 |
sig |
|
49 |
val ap_split : typ -> typ -> term -> term |
|
50 |
val factors : typ -> typ list |
|
51 |
val mk_prod : typ * typ -> typ |
|
52 |
val mk_tuple : term -> typ -> term list -> term |
|
53 |
val pseudo_type : term -> typ |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
43333
diff
changeset
|
54 |
val remove_split : Proof.context -> thm -> thm |
6049 | 55 |
val split_const : typ -> term |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
43333
diff
changeset
|
56 |
val split_rule_var : Proof.context -> term * typ * thm -> thm |
6049 | 57 |
end; |
58 |
||
59 |
||
60 |
functor CartProd_Fun (Pr: PR) : CARTPROD = |
|
61 |
struct |
|
62 |
||
15674
4a1d07bb53e2
reverted renaming of Some/None in comments and strings;
wenzelm
parents:
15570
diff
changeset
|
63 |
(* Some of these functions expect "pseudo-types" containing products, |
6049 | 64 |
as in HOL; the true ZF types would just be "i" *) |
65 |
||
66 |
fun mk_prod (T1,T2) = Type("*", [T1,T2]); |
|
67 |
||
68 |
(*Bogus product type underlying a (possibly nested) Sigma. |
|
69 |
Lets us share HOL code*) |
|
70 |
fun pseudo_type (t $ A $ Abs(_,_,B)) = |
|
26189 | 71 |
if t = Pr.sigma |
72 |
then mk_prod(pseudo_type A, pseudo_type B) |
|
69593 | 73 |
else \<^typ>\<open>i\<close> |
74 |
| pseudo_type _ = \<^typ>\<open>i\<close>; |
|
6049 | 75 |
|
76 |
(*Maps the type T1*...*Tn to [T1,...,Tn], however nested*) |
|
26189 | 77 |
fun factors (Type("*", [T1, T2])) = factors T1 @ factors T2 |
78 |
| factors T = [T]; |
|
6049 | 79 |
|
80 |
(*Make a well-typed instance of "split"*) |
|
81 |
fun split_const T = Const(Pr.split_name, |
|
82 |
[[Ind_Syntax.iT, Ind_Syntax.iT]--->T, |
|
83 |
Ind_Syntax.iT] ---> T); |
|
84 |
||
85 |
(*In ap_split S T u, term u expects separate arguments for the factors of S, |
|
86 |
with result type T. The call creates a new term expecting one argument |
|
87 |
of type S.*) |
|
88 |
fun ap_split (Type("*", [T1,T2])) T3 u = |
|
89 |
split_const T3 $ |
|
90 |
Abs("v", Ind_Syntax.iT, (*Not T1, as it involves pseudo-product types*) |
|
91 |
ap_split T2 T3 |
|
92 |
((ap_split T1 (factors T2 ---> T3) (incr_boundvars 1 u)) $ |
|
93 |
Bound 0)) |
|
94 |
| ap_split T T3 u = u; |
|
95 |
||
96 |
(*Makes a nested tuple from a list, following the product type structure*) |
|
97 |
fun mk_tuple pair (Type("*", [T1,T2])) tms = |
|
33957 | 98 |
pair $ mk_tuple pair T1 tms |
99 |
$ mk_tuple pair T2 (drop (length (factors T1)) tms) |
|
6049 | 100 |
| mk_tuple pair T (t::_) = t; |
101 |
||
102 |
(*Attempts to remove occurrences of split, and pair-valued parameters*) |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
43333
diff
changeset
|
103 |
fun remove_split ctxt = rewrite_rule ctxt [Pr.split_eq]; |
6049 | 104 |
|
105 |
(*Uncurries any Var according to its "pseudo-product type" T1 in the rule*) |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
43333
diff
changeset
|
106 |
fun split_rule_var ctxt (Var(v,_), Type("fun",[T1,T2]), rl) = |
6049 | 107 |
let val T' = factors T1 ---> T2 |
108 |
val newt = ap_split T1 T2 (Var(v,T')) |
|
109 |
in |
|
59626 | 110 |
remove_split ctxt |
111 |
(Drule.instantiate_normalize ([], |
|
60642
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
59626
diff
changeset
|
112 |
[((v, Ind_Syntax.iT-->T2), Thm.cterm_of ctxt newt)]) rl) |
6049 | 113 |
end |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
43333
diff
changeset
|
114 |
| split_rule_var _ (t,T,rl) = rl; |
6049 | 115 |
|
116 |
end; |
|
117 |