| author | wenzelm | 
| Tue, 04 Apr 2017 22:56:28 +0200 | |
| changeset 65384 | 36255c43c64c | 
| parent 64556 | 851ae0e7b09c | 
| child 67721 | 5348bea4accd | 
| permissions | -rw-r--r-- | 
| 19416 | 1  | 
(* Title: Pure/conjunction.ML  | 
2  | 
Author: Makarius  | 
|
3  | 
||
4  | 
Meta-level conjunction.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
signature CONJUNCTION =  | 
|
8  | 
sig  | 
|
9  | 
val conjunction: cterm  | 
|
10  | 
val mk_conjunction: cterm * cterm -> cterm  | 
|
| 23422 | 11  | 
val mk_conjunction_balanced: cterm list -> cterm  | 
| 19416 | 12  | 
val dest_conjunction: cterm -> cterm * cterm  | 
| 
30823
 
eb99b9134f2e
added dest_conjunctions (cf. Logic.dest_conjunctions);
 
wenzelm 
parents: 
29606 
diff
changeset
 | 
13  | 
val dest_conjunctions: cterm -> cterm list  | 
| 19416 | 14  | 
val cong: thm -> thm -> thm  | 
| 23422 | 15  | 
val convs: (cterm -> thm) -> cterm -> thm  | 
| 19416 | 16  | 
val conjunctionD1: thm  | 
17  | 
val conjunctionD2: thm  | 
|
18  | 
val conjunctionI: thm  | 
|
19  | 
val intr: thm -> thm -> thm  | 
|
| 23422 | 20  | 
val intr_balanced: thm list -> thm  | 
| 19416 | 21  | 
val elim: thm -> thm * thm  | 
| 60623 | 22  | 
val elim_conjunctions: thm -> thm list  | 
| 23422 | 23  | 
val elim_balanced: int -> thm -> thm list  | 
24  | 
val curry_balanced: int -> thm -> thm  | 
|
25  | 
val uncurry_balanced: int -> thm -> thm  | 
|
| 19416 | 26  | 
end;  | 
27  | 
||
28  | 
structure Conjunction: CONJUNCTION =  | 
|
29  | 
struct  | 
|
30  | 
||
31  | 
(** abstract syntax **)  | 
|
32  | 
||
| 62876 | 33  | 
fun certify t = Thm.global_cterm_of (Context.the_global_context ()) t;  | 
| 33384 | 34  | 
val read_prop = certify o Simple_Syntax.read_prop;  | 
| 19416 | 35  | 
|
| 26485 | 36  | 
val true_prop = certify Logic.true_prop;  | 
37  | 
val conjunction = certify Logic.conjunction;  | 
|
| 23422 | 38  | 
|
| 
46497
 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 
wenzelm 
parents: 
43329 
diff
changeset
 | 
39  | 
fun mk_conjunction (A, B) = Thm.apply (Thm.apply conjunction A) B;  | 
| 19416 | 40  | 
|
| 23422 | 41  | 
fun mk_conjunction_balanced [] = true_prop  | 
| 32765 | 42  | 
| mk_conjunction_balanced ts = Balanced_Tree.make mk_conjunction ts;  | 
| 23422 | 43  | 
|
| 19416 | 44  | 
fun dest_conjunction ct =  | 
45  | 
(case Thm.term_of ct of  | 
|
| 26424 | 46  | 
    (Const ("Pure.conjunction", _) $ _ $ _) => Thm.dest_binop ct
 | 
| 23422 | 47  | 
  | _ => raise TERM ("dest_conjunction", [Thm.term_of ct]));
 | 
| 19416 | 48  | 
|
| 
30823
 
eb99b9134f2e
added dest_conjunctions (cf. Logic.dest_conjunctions);
 
wenzelm 
parents: 
29606 
diff
changeset
 | 
49  | 
fun dest_conjunctions ct =  | 
| 
 
eb99b9134f2e
added dest_conjunctions (cf. Logic.dest_conjunctions);
 
wenzelm 
parents: 
29606 
diff
changeset
 | 
50  | 
(case try dest_conjunction ct of  | 
| 
 
eb99b9134f2e
added dest_conjunctions (cf. Logic.dest_conjunctions);
 
wenzelm 
parents: 
29606 
diff
changeset
 | 
51  | 
NONE => [ct]  | 
| 
 
eb99b9134f2e
added dest_conjunctions (cf. Logic.dest_conjunctions);
 
wenzelm 
parents: 
29606 
diff
changeset
 | 
52  | 
| SOME (A, B) => dest_conjunctions A @ dest_conjunctions B);  | 
| 
 
eb99b9134f2e
added dest_conjunctions (cf. Logic.dest_conjunctions);
 
wenzelm 
parents: 
29606 
diff
changeset
 | 
53  | 
|
| 19416 | 54  | 
|
55  | 
||
56  | 
(** derived rules **)  | 
|
57  | 
||
58  | 
(* conversion *)  | 
|
59  | 
||
60  | 
val cong = Thm.combination o Thm.combination (Thm.reflexive conjunction);  | 
|
61  | 
||
| 23422 | 62  | 
fun convs cv ct =  | 
63  | 
(case try dest_conjunction ct of  | 
|
64  | 
NONE => cv ct  | 
|
65  | 
| SOME (A, B) => cong (convs cv A) (convs cv B));  | 
|
| 19416 | 66  | 
|
67  | 
||
68  | 
(* intro/elim *)  | 
|
69  | 
||
70  | 
local  | 
|
71  | 
||
| 
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: 
60623 
diff
changeset
 | 
72  | 
val A = read_prop "A" and vA = (("A", 0), propT);
 | 
| 
 
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: 
60623 
diff
changeset
 | 
73  | 
val B = read_prop "B" and vB = (("B", 0), propT);
 | 
| 24241 | 74  | 
val C = read_prop "C";  | 
75  | 
val ABC = read_prop "A ==> B ==> C";  | 
|
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28674 
diff
changeset
 | 
76  | 
val A_B = read_prop "A &&& B";  | 
| 19416 | 77  | 
|
| 26424 | 78  | 
val conjunction_def =  | 
| 62876 | 79  | 
Thm.unvarify_axiom (Context.the_global_context ()) "Pure.conjunction_def";  | 
| 19416 | 80  | 
|
81  | 
fun conjunctionD which =  | 
|
82  | 
Drule.implies_intr_list [A, B] (Thm.assume (which (A, B))) COMP  | 
|
| 26653 | 83  | 
Thm.forall_elim_vars 0 (Thm.equal_elim conjunction_def (Thm.assume A_B));  | 
| 19416 | 84  | 
|
85  | 
in  | 
|
86  | 
||
| 56436 | 87  | 
val conjunctionD1 =  | 
| 64556 | 88  | 
  Drule.store_standard_thm (Binding.make ("conjunctionD1", \<^here>)) (conjunctionD #1);
 | 
| 56436 | 89  | 
|
90  | 
val conjunctionD2 =  | 
|
| 64556 | 91  | 
  Drule.store_standard_thm (Binding.make ("conjunctionD2", \<^here>)) (conjunctionD #2);
 | 
| 19416 | 92  | 
|
| 33277 | 93  | 
val conjunctionI =  | 
| 64556 | 94  | 
  Drule.store_standard_thm (Binding.make ("conjunctionI", \<^here>))
 | 
| 33277 | 95  | 
(Drule.implies_intr_list [A, B]  | 
96  | 
(Thm.equal_elim  | 
|
97  | 
(Thm.symmetric conjunction_def)  | 
|
98  | 
(Thm.forall_intr C (Thm.implies_intr ABC  | 
|
99  | 
(Drule.implies_elim_list (Thm.assume ABC) [Thm.assume A, Thm.assume B])))));  | 
|
| 19416 | 100  | 
|
| 23422 | 101  | 
|
| 
20508
 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 
wenzelm 
parents: 
20260 
diff
changeset
 | 
102  | 
fun intr tha thb =  | 
| 
 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 
wenzelm 
parents: 
20260 
diff
changeset
 | 
103  | 
Thm.implies_elim  | 
| 
 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 
wenzelm 
parents: 
20260 
diff
changeset
 | 
104  | 
(Thm.implies_elim  | 
| 
 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 
wenzelm 
parents: 
20260 
diff
changeset
 | 
105  | 
(Thm.instantiate ([], [(vA, Thm.cprop_of tha), (vB, Thm.cprop_of thb)]) conjunctionI)  | 
| 
 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 
wenzelm 
parents: 
20260 
diff
changeset
 | 
106  | 
tha)  | 
| 
 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 
wenzelm 
parents: 
20260 
diff
changeset
 | 
107  | 
thb;  | 
| 19416 | 108  | 
|
109  | 
fun elim th =  | 
|
| 
20508
 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 
wenzelm 
parents: 
20260 
diff
changeset
 | 
110  | 
let  | 
| 
 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 
wenzelm 
parents: 
20260 
diff
changeset
 | 
111  | 
val (A, B) = dest_conjunction (Thm.cprop_of th)  | 
| 
 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 
wenzelm 
parents: 
20260 
diff
changeset
 | 
112  | 
handle TERM (msg, _) => raise THM (msg, 0, [th]);  | 
| 
 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 
wenzelm 
parents: 
20260 
diff
changeset
 | 
113  | 
val inst = Thm.instantiate ([], [(vA, A), (vB, B)]);  | 
| 
 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 
wenzelm 
parents: 
20260 
diff
changeset
 | 
114  | 
in  | 
| 
 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 
wenzelm 
parents: 
20260 
diff
changeset
 | 
115  | 
(Thm.implies_elim (inst conjunctionD1) th,  | 
| 
 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 
wenzelm 
parents: 
20260 
diff
changeset
 | 
116  | 
Thm.implies_elim (inst conjunctionD2) th)  | 
| 
 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 
wenzelm 
parents: 
20260 
diff
changeset
 | 
117  | 
end;  | 
| 19416 | 118  | 
|
| 23422 | 119  | 
end;  | 
120  | 
||
| 60623 | 121  | 
fun elim_conjunctions th =  | 
122  | 
(case try elim th of  | 
|
123  | 
NONE => [th]  | 
|
124  | 
| SOME (th1, th2) => elim_conjunctions th1 @ elim_conjunctions th2);  | 
|
125  | 
||
| 23422 | 126  | 
|
| 
23535
 
58147e5bd070
removed obsolete mk_conjunction_list, intr/elim_list;
 
wenzelm 
parents: 
23422 
diff
changeset
 | 
127  | 
(* balanced conjuncts *)  | 
| 23422 | 128  | 
|
129  | 
fun intr_balanced [] = asm_rl  | 
|
| 32765 | 130  | 
| intr_balanced ths = Balanced_Tree.make (uncurry intr) ths;  | 
| 23422 | 131  | 
|
132  | 
fun elim_balanced 0 _ = []  | 
|
| 32765 | 133  | 
| elim_balanced n th = Balanced_Tree.dest elim n th;  | 
| 19416 | 134  | 
|
135  | 
||
136  | 
(* currying *)  | 
|
137  | 
||
138  | 
local  | 
|
139  | 
||
| 62876 | 140  | 
val bootstrap_thy = Context.the_global_context ();  | 
| 60815 | 141  | 
|
142  | 
fun conjs n =  | 
|
143  | 
let  | 
|
144  | 
val As =  | 
|
145  | 
map (fn A => Thm.global_cterm_of bootstrap_thy (Free (A, propT)))  | 
|
146  | 
(Name.invent Name.context "A" n);  | 
|
| 23422 | 147  | 
in (As, mk_conjunction_balanced As) end;  | 
| 19416 | 148  | 
|
| 24241 | 149  | 
val B = read_prop "B";  | 
| 19416 | 150  | 
|
151  | 
fun comp_rule th rule =  | 
|
| 20260 | 152  | 
Thm.adjust_maxidx_thm ~1 (th COMP  | 
| 
35985
 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 
wenzelm 
parents: 
35845 
diff
changeset
 | 
153  | 
(rule |> Thm.forall_intr_frees |> Thm.forall_elim_vars (Thm.maxidx_of th + 1)));  | 
| 19416 | 154  | 
|
155  | 
in  | 
|
156  | 
||
157  | 
(*  | 
|
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28674 
diff
changeset
 | 
158  | 
A1 &&& ... &&& An ==> B  | 
| 19416 | 159  | 
-----------------------  | 
160  | 
A1 ==> ... ==> An ==> B  | 
|
161  | 
*)  | 
|
| 23422 | 162  | 
fun curry_balanced n th =  | 
163  | 
if n < 2 then th  | 
|
164  | 
else  | 
|
165  | 
let  | 
|
| 60815 | 166  | 
val (As, C) = conjs n;  | 
| 23422 | 167  | 
val D = Drule.mk_implies (C, B);  | 
168  | 
in  | 
|
169  | 
comp_rule th  | 
|
170  | 
(Thm.implies_elim (Thm.assume D) (intr_balanced (map Thm.assume As))  | 
|
171  | 
|> Drule.implies_intr_list (D :: As))  | 
|
172  | 
end;  | 
|
| 19416 | 173  | 
|
174  | 
(*  | 
|
175  | 
A1 ==> ... ==> An ==> B  | 
|
176  | 
-----------------------  | 
|
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28674 
diff
changeset
 | 
177  | 
A1 &&& ... &&& An ==> B  | 
| 19416 | 178  | 
*)  | 
| 23422 | 179  | 
fun uncurry_balanced n th =  | 
180  | 
if n < 2 then th  | 
|
181  | 
else  | 
|
182  | 
let  | 
|
| 60815 | 183  | 
val (As, C) = conjs n;  | 
| 23422 | 184  | 
val D = Drule.list_implies (As, B);  | 
185  | 
in  | 
|
186  | 
comp_rule th  | 
|
187  | 
(Drule.implies_elim_list (Thm.assume D) (elim_balanced n (Thm.assume C))  | 
|
188  | 
|> Drule.implies_intr_list [D, C])  | 
|
189  | 
end;  | 
|
| 19416 | 190  | 
|
191  | 
end;  | 
|
192  | 
||
193  | 
end;  |