added dest_conjunctions (cf. Logic.dest_conjunctions);
1 (* Title: Pure/conjunction.ML
4 Meta-level conjunction.
7 signature CONJUNCTION =
10 val mk_conjunction: cterm * cterm -> cterm
11 val mk_conjunction_balanced: cterm list -> cterm
12 val dest_conjunction: cterm -> cterm * cterm
13 val dest_conjunctions: cterm -> cterm list
14 val cong: thm -> thm -> thm
15 val convs: (cterm -> thm) -> cterm -> thm
16 val conjunctionD1: thm
17 val conjunctionD2: thm
19 val intr: thm -> thm -> thm
20 val intr_balanced: thm list -> thm
21 val elim: thm -> thm * thm
22 val elim_balanced: int -> thm -> thm list
23 val curry_balanced: int -> thm -> thm
24 val uncurry_balanced: int -> thm -> thm
27 structure Conjunction: CONJUNCTION =
30 (** abstract syntax **)
32 fun certify t = Thm.cterm_of (Context.the_theory (Context.the_thread_data ())) t;
33 val read_prop = certify o SimpleSyntax.read_prop;
35 val true_prop = certify Logic.true_prop;
36 val conjunction = certify Logic.conjunction;
38 fun mk_conjunction (A, B) = Thm.capply (Thm.capply conjunction A) B;
40 fun mk_conjunction_balanced [] = true_prop
41 | mk_conjunction_balanced ts = BalancedTree.make mk_conjunction ts;
43 fun dest_conjunction ct =
44 (case Thm.term_of ct of
45 (Const ("Pure.conjunction", _) $ _ $ _) => Thm.dest_binop ct
46 | _ => raise TERM ("dest_conjunction", [Thm.term_of ct]));
48 fun dest_conjunctions ct =
49 (case try dest_conjunction ct of
51 | SOME (A, B) => dest_conjunctions A @ dest_conjunctions B);
59 val cong = Thm.combination o Thm.combination (Thm.reflexive conjunction);
62 (case try dest_conjunction ct of
64 | SOME (A, B) => cong (convs cv A) (convs cv B));
71 val A = read_prop "A" and vA = read_prop "?A";
72 val B = read_prop "B" and vB = read_prop "?B";
73 val C = read_prop "C";
74 val ABC = read_prop "A ==> B ==> C";
75 val A_B = read_prop "A &&& B";
78 Thm.unvarify (Thm.axiom (Context.the_theory (Context.the_thread_data ())) "Pure.conjunction_def");
80 fun conjunctionD which =
81 Drule.implies_intr_list [A, B] (Thm.assume (which (A, B))) COMP
82 Thm.forall_elim_vars 0 (Thm.equal_elim conjunction_def (Thm.assume A_B));
86 val conjunctionD1 = Drule.store_standard_thm "conjunctionD1" (conjunctionD #1);
87 val conjunctionD2 = Drule.store_standard_thm "conjunctionD2" (conjunctionD #2);
89 val conjunctionI = Drule.store_standard_thm "conjunctionI"
90 (Drule.implies_intr_list [A, B]
92 (Thm.symmetric conjunction_def)
93 (Thm.forall_intr C (Thm.implies_intr ABC
94 (Drule.implies_elim_list (Thm.assume ABC) [Thm.assume A, Thm.assume B])))));
100 (Thm.instantiate ([], [(vA, Thm.cprop_of tha), (vB, Thm.cprop_of thb)]) conjunctionI)
106 val (A, B) = dest_conjunction (Thm.cprop_of th)
107 handle TERM (msg, _) => raise THM (msg, 0, [th]);
108 val inst = Thm.instantiate ([], [(vA, A), (vB, B)]);
110 (Thm.implies_elim (inst conjunctionD1) th,
111 Thm.implies_elim (inst conjunctionD2) th)
117 (* balanced conjuncts *)
119 fun intr_balanced [] = asm_rl
120 | intr_balanced ths = BalancedTree.make (uncurry intr) ths;
122 fun elim_balanced 0 _ = []
123 | elim_balanced n th = BalancedTree.dest elim n th;
131 let val As = map (fn A => Thm.cterm_of thy (Free (A, propT))) (Name.invents Name.context "A" n)
132 in (As, mk_conjunction_balanced As) end;
134 val B = read_prop "B";
136 fun comp_rule th rule =
137 Thm.adjust_maxidx_thm ~1 (th COMP
138 (rule |> Drule.forall_intr_frees |> Thm.forall_elim_vars (Thm.maxidx_of th + 1)));
143 A1 &&& ... &&& An ==> B
144 -----------------------
145 A1 ==> ... ==> An ==> B
147 fun curry_balanced n th =
151 val thy = Thm.theory_of_thm th;
152 val (As, C) = conjs thy n;
153 val D = Drule.mk_implies (C, B);
156 (Thm.implies_elim (Thm.assume D) (intr_balanced (map Thm.assume As))
157 |> Drule.implies_intr_list (D :: As))
161 A1 ==> ... ==> An ==> B
162 -----------------------
163 A1 &&& ... &&& An ==> B
165 fun uncurry_balanced n th =
169 val thy = Thm.theory_of_thm th;
170 val (As, C) = conjs thy n;
171 val D = Drule.list_implies (As, B);
174 (Drule.implies_elim_list (Thm.assume D) (elim_balanced n (Thm.assume C))
175 |> Drule.implies_intr_list [D, C])