| author | wenzelm | 
| Wed, 01 May 2024 20:26:20 +0200 | |
| changeset 80166 | 825f35bae74b | 
| parent 79339 | 8eb7e325f40d | 
| child 81507 | 08574da77b4a | 
| 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: 
29606diff
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: 
43329diff
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: 
29606diff
changeset | 49 | fun dest_conjunctions ct = | 
| 
eb99b9134f2e
added dest_conjunctions (cf. Logic.dest_conjunctions);
 wenzelm parents: 
29606diff
changeset | 50 | (case try dest_conjunction ct of | 
| 
eb99b9134f2e
added dest_conjunctions (cf. Logic.dest_conjunctions);
 wenzelm parents: 
29606diff
changeset | 51 | NONE => [ct] | 
| 
eb99b9134f2e
added dest_conjunctions (cf. Logic.dest_conjunctions);
 wenzelm parents: 
29606diff
changeset | 52 | | SOME (A, B) => dest_conjunctions A @ dest_conjunctions B); | 
| 
eb99b9134f2e
added dest_conjunctions (cf. Logic.dest_conjunctions);
 wenzelm parents: 
29606diff
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: 
60623diff
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: 
60623diff
changeset | 73 | val B = read_prop "B" and vB = (("B", 0), propT);
 | 
| 24241 | 74 | val C = read_prop "C"; | 
| 67721 | 75 | val ABC = read_prop "A \<Longrightarrow> B \<Longrightarrow> C"; | 
| 28856 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
28674diff
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: 
20260diff
changeset | 102 | fun intr tha thb = | 
| 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 wenzelm parents: 
20260diff
changeset | 103 | Thm.implies_elim | 
| 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 wenzelm parents: 
20260diff
changeset | 104 | (Thm.implies_elim | 
| 77879 | 105 | (Thm.instantiate (TVars.empty, Vars.make2 (vA, Thm.cprop_of tha) (vB, Thm.cprop_of thb)) | 
| 74282 | 106 | conjunctionI) | 
| 20508 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 wenzelm parents: 
20260diff
changeset | 107 | tha) | 
| 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 wenzelm parents: 
20260diff
changeset | 108 | thb; | 
| 19416 | 109 | |
| 110 | fun elim th = | |
| 20508 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 wenzelm parents: 
20260diff
changeset | 111 | let | 
| 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 wenzelm parents: 
20260diff
changeset | 112 | val (A, B) = dest_conjunction (Thm.cprop_of th) | 
| 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 wenzelm parents: 
20260diff
changeset | 113 | handle TERM (msg, _) => raise THM (msg, 0, [th]); | 
| 77879 | 114 | val inst = Thm.instantiate (TVars.empty, Vars.make2 (vA, A) (vB, B)); | 
| 20508 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 wenzelm parents: 
20260diff
changeset | 115 | in | 
| 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 wenzelm parents: 
20260diff
changeset | 116 | (Thm.implies_elim (inst conjunctionD1) th, | 
| 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 wenzelm parents: 
20260diff
changeset | 117 | Thm.implies_elim (inst conjunctionD2) th) | 
| 
8182d961c7cc
intr/elim: use constant complexity thanks to tuned Thm.instantiate/implies_elim;
 wenzelm parents: 
20260diff
changeset | 118 | end; | 
| 19416 | 119 | |
| 23422 | 120 | end; | 
| 121 | ||
| 60623 | 122 | fun elim_conjunctions th = | 
| 123 | (case try elim th of | |
| 124 | NONE => [th] | |
| 125 | | SOME (th1, th2) => elim_conjunctions th1 @ elim_conjunctions th2); | |
| 126 | ||
| 23422 | 127 | |
| 23535 
58147e5bd070
removed obsolete mk_conjunction_list, intr/elim_list;
 wenzelm parents: 
23422diff
changeset | 128 | (* balanced conjuncts *) | 
| 23422 | 129 | |
| 130 | fun intr_balanced [] = asm_rl | |
| 32765 | 131 | | intr_balanced ths = Balanced_Tree.make (uncurry intr) ths; | 
| 23422 | 132 | |
| 133 | fun elim_balanced 0 _ = [] | |
| 32765 | 134 | | elim_balanced n th = Balanced_Tree.dest elim n th; | 
| 19416 | 135 | |
| 136 | ||
| 137 | (* currying *) | |
| 138 | ||
| 139 | local | |
| 140 | ||
| 62876 | 141 | val bootstrap_thy = Context.the_global_context (); | 
| 60815 | 142 | |
| 143 | fun conjs n = | |
| 144 | let | |
| 145 | val As = | |
| 146 | map (fn A => Thm.global_cterm_of bootstrap_thy (Free (A, propT))) | |
| 79339 | 147 | (Name.invent Name.context "" n); | 
| 23422 | 148 | in (As, mk_conjunction_balanced As) end; | 
| 19416 | 149 | |
| 24241 | 150 | val B = read_prop "B"; | 
| 19416 | 151 | |
| 79338 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 152 | fun gen_rule idx rule = | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 153 |   let val frees = Names.build (rule |> Thm.fold_terms {hyps = true} Names.add_free_names)
 | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 154 | in rule |> Thm.generalize (Names.empty, frees) idx end; | 
| 19416 | 155 | |
| 156 | (* | |
| 67721 | 157 | A1 &&& ... &&& An \<Longrightarrow> B | 
| 19416 | 158 | ----------------------- | 
| 67721 | 159 | A1 \<Longrightarrow> ... \<Longrightarrow> An \<Longrightarrow> B | 
| 19416 | 160 | *) | 
| 79338 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 161 | fun curry_balanced_rule idx n = | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 162 | let | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 163 | val (As, C) = conjs n; | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 164 | val D = Drule.mk_implies (C, B); | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 165 | in | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 166 | Thm.implies_elim (Thm.assume D) (intr_balanced (map Thm.assume As)) | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 167 | |> Drule.implies_intr_list (D :: As) | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 168 | |> gen_rule idx | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 169 | end; | 
| 19416 | 170 | |
| 171 | (* | |
| 67721 | 172 | A1 \<Longrightarrow> ... \<Longrightarrow> An \<Longrightarrow> B | 
| 19416 | 173 | ----------------------- | 
| 67721 | 174 | A1 &&& ... &&& An \<Longrightarrow> B | 
| 19416 | 175 | *) | 
| 79338 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 176 | fun uncurry_balanced_rule idx n = | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 177 | let | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 178 | val (As, C) = conjs n; | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 179 | val D = Drule.list_implies (As, B); | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 180 | in | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 181 | Drule.implies_elim_list (Thm.assume D) (elim_balanced n (Thm.assume C)) | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 182 | |> Drule.implies_intr_list [D, C] | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 183 | |> gen_rule idx | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 184 | end; | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 185 | |
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 186 | |
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 187 | (* static vs. dynamic rules *) | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 188 | |
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 189 | fun make_rules make = (make, Vector.tabulate (10, fn i => make 0 (i + 2))); | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 190 | |
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 191 | fun apply_rule (make, rules) n thm = | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 192 | if n < 2 then thm | 
| 23422 | 193 | else | 
| 194 | let | |
| 79338 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 195 | val idx = Thm.maxidx_of thm + 1; | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 196 | val rule = | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 197 | (case try Vector.sub (rules, n - 2) of | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 198 | SOME rule => Thm.incr_indexes idx rule | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 199 | | NONE => make idx n); | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 200 | in Thm.adjust_maxidx_thm ~1 (thm COMP rule) end; | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 201 | |
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 202 | in | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 203 | |
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 204 | val curry_balanced = apply_rule (make_rules curry_balanced_rule); | 
| 
b3b0950ef24e
minor performance tuning: static vs. dynamic rules;
 wenzelm parents: 
79337diff
changeset | 205 | val uncurry_balanced = apply_rule (make_rules uncurry_balanced_rule); | 
| 19416 | 206 | |
| 207 | end; | |
| 208 | ||
| 209 | end; |