| author | wenzelm | 
| Thu, 10 Jul 2025 15:08:26 +0200 | |
| changeset 82837 | cd566dbe9f48 | 
| parent 81954 | 6f2bcdfa9a19 | 
| permissions | -rw-r--r-- | 
| 23175 | 1 | (* Title: Tools/IsaPlanner/isand.ML | 
| 23171 | 2 | Author: Lucas Dixon, University of Edinburgh | 
| 23175 | 3 | |
| 52244 | 4 | Natural Deduction tools (obsolete). | 
| 23171 | 5 | |
| 23175 | 6 | For working with Isabelle theorems in a natural detuction style. | 
| 7 | ie, not having to deal with meta level quantified varaibles, | |
| 8 | instead, we work with newly introduced frees, and hide the | |
| 9 | "all"'s, exporting results from theorems proved with the frees, to | |
| 10 | solve the all cases of the previous goal. This allows resolution | |
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
46777diff
changeset | 11 | to do proof search normally. | 
| 23171 | 12 | |
| 23175 | 13 | Note: A nice idea: allow exporting to solve any subgoal, thus | 
| 14 | allowing the interleaving of proof, or provide a structure for the | |
| 15 | ordering of proof, thus allowing proof attempts in parrell, but | |
| 16 | recording the order to apply things in. | |
| 23171 | 17 | |
| 23175 | 18 | THINK: are we really ok with our varify name w.r.t the prop - do | 
| 19 | we also need to avoid names in the hidden hyps? What about | |
| 20 | unification contraints in flex-flex pairs - might they also have | |
| 21 | extra free vars? | |
| 23171 | 22 | *) | 
| 23 | ||
| 24 | signature ISA_ND = | |
| 25 | sig | |
| 52244 | 26 | val variant_names: Proof.context -> term list -> string list -> string list | 
| 49340 
25fc6e0da459
observe context more carefully when producing "fresh" variables -- for increased chances that method "subst" works in local context (including that of forked proofs);
 wenzelm parents: 
49339diff
changeset | 27 | |
| 23171 | 28 | (* meta level fixed params (i.e. !! vars) *) | 
| 52244 | 29 | val fix_alls_term: Proof.context -> int -> term -> term * term list | 
| 23171 | 30 | |
| 31 | (* assumptions/subgoals *) | |
| 52244 | 32 | val fixed_subgoal_thms: Proof.context -> thm -> thm list * (thm list -> thm) | 
| 23171 | 33 | end | 
| 34 | ||
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
46777diff
changeset | 35 | structure IsaND : ISA_ND = | 
| 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
46777diff
changeset | 36 | struct | 
| 23171 | 37 | |
| 38 | (* datatype to capture an exported result, ie a fix or assume. *) | |
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
46777diff
changeset | 39 | datatype export = | 
| 52244 | 40 | Export of | 
| 41 |    {fixes : Thm.cterm list, (* fixed vars *)
 | |
| 42 | assumes : Thm.cterm list, (* hidden hyps/assumed prems *) | |
| 43 | sgid : int, | |
| 44 | gth : Thm.thm}; (* subgoal/goalthm *) | |
| 23171 | 45 | |
| 46 | (* exporting function that takes a solution to the fixed/assumed goal, | |
| 47 | and uses this to solve the subgoal in the main theorem *) | |
| 52244 | 48 | fun export_solution (Export {fixes = cfvs, assumes = hcprems, sgid = i, gth = gth}) solth =
 | 
| 49 | let | |
| 50 | val solth' = solth | |
| 51 | |> Drule.implies_intr_list hcprems | |
| 52 | |> Drule.forall_intr_list cfvs; | |
| 52467 | 53 | in Drule.compose (solth', i, gth) end; | 
| 23171 | 54 | |
| 49340 
25fc6e0da459
observe context more carefully when producing "fresh" variables -- for increased chances that method "subst" works in local context (including that of forked proofs);
 wenzelm parents: 
49339diff
changeset | 55 | fun variant_names ctxt ts xs = | 
| 
25fc6e0da459
observe context more carefully when producing "fresh" variables -- for increased chances that method "subst" works in local context (including that of forked proofs);
 wenzelm parents: 
49339diff
changeset | 56 | let | 
| 
25fc6e0da459
observe context more carefully when producing "fresh" variables -- for increased chances that method "subst" works in local context (including that of forked proofs);
 wenzelm parents: 
49339diff
changeset | 57 | val names = | 
| 
25fc6e0da459
observe context more carefully when producing "fresh" variables -- for increased chances that method "subst" works in local context (including that of forked proofs);
 wenzelm parents: 
49339diff
changeset | 58 | Variable.names_of ctxt | 
| 81515 | 59 | |> fold (fn t => Term.declare_free_names t #> Term.declare_var_names (K true) t) ts; | 
| 81521 | 60 | in Name.variants names xs end; | 
| 23171 | 61 | |
| 62 | (* fix parameters of a subgoal "i", as free variables, and create an | |
| 63 | exporting function that will use the result of this proved goal to | |
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
46777diff
changeset | 64 | show the goal in the original theorem. | 
| 23171 | 65 | |
| 66 | Note, an advantage of this over Isar is that it supports instantiation | |
| 67 | of unkowns in the earlier theorem, ie we can do instantiation of meta | |
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
46777diff
changeset | 68 | vars! | 
| 23171 | 69 | |
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
46777diff
changeset | 70 | avoids constant, free and vars names. | 
| 23171 | 71 | |
| 72 | loosely corresponds to: | |
| 73 | Given "[| SG0; ... !! x. As ==> SGi x; ... SGm |] ==> G" : thm | |
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
46777diff
changeset | 74 | Result: | 
| 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
46777diff
changeset | 75 |   ("(As ==> SGi x') ==> (As ==> SGi x')" : thm,
 | 
| 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
46777diff
changeset | 76 | expf : | 
| 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
46777diff
changeset | 77 |      ("As ==> SGi x'" : thm) ->
 | 
| 23171 | 78 |      ("[| SG0; ... SGi-1; SGi+1; ... SGm |] ==> G") : thm)
 | 
| 79 | *) | |
| 49340 
25fc6e0da459
observe context more carefully when producing "fresh" variables -- for increased chances that method "subst" works in local context (including that of forked proofs);
 wenzelm parents: 
49339diff
changeset | 80 | fun fix_alls_term ctxt i t = | 
| 52244 | 81 | let | 
| 82 | val gt = Logic.get_goal t i; | |
| 83 | val body = Term.strip_all_body gt; | |
| 84 | val alls = rev (Term.strip_all_vars gt); | |
| 85 | val xs = variant_names ctxt [t] (map fst alls); | |
| 86 | val fvs = map Free (xs ~~ map snd alls); | |
| 87 | in ((subst_bounds (fvs,body)), fvs) end; | |
| 23171 | 88 | |
| 49340 
25fc6e0da459
observe context more carefully when producing "fresh" variables -- for increased chances that method "subst" works in local context (including that of forked proofs);
 wenzelm parents: 
49339diff
changeset | 89 | fun fix_alls_cterm ctxt i th = | 
| 52244 | 90 | let | 
| 91 | val (fixedbody, fvs) = fix_alls_term ctxt i (Thm.prop_of th); | |
| 59641 | 92 | val cfvs = rev (map (Thm.cterm_of ctxt) fvs); | 
| 93 | val ct_body = Thm.cterm_of ctxt fixedbody; | |
| 52244 | 94 | in (ct_body, cfvs) end; | 
| 23171 | 95 | |
| 49340 
25fc6e0da459
observe context more carefully when producing "fresh" variables -- for increased chances that method "subst" works in local context (including that of forked proofs);
 wenzelm parents: 
49339diff
changeset | 96 | fun fix_alls' ctxt i = apfst Thm.trivial o fix_alls_cterm ctxt i; | 
| 23171 | 97 | |
| 98 | ||
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
46777diff
changeset | 99 | (* hide other goals *) | 
| 23171 | 100 | (* note the export goal is rotated by (i - 1) and will have to be | 
| 101 | unrotated to get backto the originial position(s) *) | |
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
46777diff
changeset | 102 | fun hide_other_goals th = | 
| 52244 | 103 | let | 
| 104 | (* tl beacuse fst sg is the goal we are interested in *) | |
| 81954 
6f2bcdfa9a19
misc tuning: more concise operations on prems (without change of exceptions);
 wenzelm parents: 
81521diff
changeset | 105 | val cprems = tl (Thm.cprems_of th); | 
| 52244 | 106 | val aprems = map Thm.assume cprems; | 
| 107 | in (Drule.implies_elim_list (Drule.rotate_prems 1 th) aprems, cprems) end; | |
| 23171 | 108 | |
| 109 | (* a nicer version of the above that leaves only a single subgoal (the | |
| 110 | other subgoals are hidden hyps, that the exporter suffles about) | |
| 111 | namely the subgoal that we were trying to solve. *) | |
| 112 | (* loosely corresponds to: | |
| 113 | Given "[| SG0; ... !! x. As ==> SGi x; ... SGm |] ==> G" : thm | |
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
46777diff
changeset | 114 | Result: | 
| 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
46777diff
changeset | 115 |   ("(As ==> SGi x') ==> SGi x'" : thm,
 | 
| 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
46777diff
changeset | 116 | expf : | 
| 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
46777diff
changeset | 117 |      ("SGi x'" : thm) ->
 | 
| 23171 | 118 |      ("[| SG0; ... SGi-1; SGi+1; ... SGm |] ==> G") : thm)
 | 
| 119 | *) | |
| 49340 
25fc6e0da459
observe context more carefully when producing "fresh" variables -- for increased chances that method "subst" works in local context (including that of forked proofs);
 wenzelm parents: 
49339diff
changeset | 120 | fun fix_alls ctxt i th = | 
| 52244 | 121 | let | 
| 122 | val (fixed_gth, fixedvars) = fix_alls' ctxt i th | |
| 123 | val (sml_gth, othergoals) = hide_other_goals fixed_gth | |
| 124 |   in (sml_gth, Export {fixes = fixedvars, assumes = othergoals, sgid = i, gth = th}) end;
 | |
| 23171 | 125 | |
| 126 | ||
| 127 | (* Fixme: allow different order of subgoals given to expf *) | |
| 128 | (* make each subgoal into a separate thm that needs to be proved *) | |
| 129 | (* loosely corresponds to: | |
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
46777diff
changeset | 130 | Given | 
| 23171 | 131 | "[| SG0; ... SGm |] ==> G" : thm | 
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
46777diff
changeset | 132 | Result: | 
| 23171 | 133 | (["SG0 ==> SG0", ... ,"SGm ==> SGm"] : thm list, -- goals | 
| 134 | ["SG0", ..., "SGm"] : thm list -> -- export function | |
| 135 | "G" : thm) | |
| 136 | *) | |
| 60358 | 137 | fun subgoal_thms ctxt th = | 
| 52244 | 138 | let | 
| 59582 | 139 | val t = Thm.prop_of th; | 
| 23171 | 140 | |
| 52244 | 141 | val prems = Logic.strip_imp_prems t; | 
| 60358 | 142 | val aprems = map (Thm.trivial o Thm.cterm_of ctxt) prems; | 
| 23171 | 143 | |
| 52244 | 144 | fun explortf premths = Drule.implies_elim_list th premths; | 
| 145 | in (aprems, explortf) end; | |
| 23171 | 146 | |
| 147 | ||
| 148 | (* Fixme: allow different order of subgoals in exportf *) | |
| 149 | (* as above, but also fix all parameters in all subgoals, and uses | |
| 150 | fix_alls, not fix_alls', ie doesn't leave extra asumptions as apparent | |
| 151 | subgoals. *) | |
| 152 | (* loosely corresponds to: | |
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
46777diff
changeset | 153 | Given | 
| 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
46777diff
changeset | 154 | "[| !! x0s. A0s x0s ==> SG0 x0s; | 
| 23171 | 155 | ...; !! xms. Ams xms ==> SGm xms|] ==> G" : thm | 
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
46777diff
changeset | 156 | Result: | 
| 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
46777diff
changeset | 157 | (["(A0s x0s' ==> SG0 x0s') ==> SG0 x0s'", | 
| 23171 | 158 | ... ,"(Ams xms' ==> SGm xms') ==> SGm xms'"] : thm list, -- goals | 
| 159 | ["SG0 x0s'", ..., "SGm xms'"] : thm list -> -- export function | |
| 160 | "G" : thm) | |
| 161 | *) | |
| 162 | (* requires being given solutions! *) | |
| 49340 
25fc6e0da459
observe context more carefully when producing "fresh" variables -- for increased chances that method "subst" works in local context (including that of forked proofs);
 wenzelm parents: 
49339diff
changeset | 163 | fun fixed_subgoal_thms ctxt th = | 
| 52244 | 164 | let | 
| 60358 | 165 | val (subgoals, expf) = subgoal_thms ctxt th; | 
| 52244 | 166 | (* fun export_sg (th, exp) = exp th; *) | 
| 167 | fun export_sgs expfs solthms = | |
| 168 | expf (map2 (curry (op |>)) solthms expfs); | |
| 169 | (* expf (map export_sg (ths ~~ expfs)); *) | |
| 170 | in | |
| 171 | apsnd export_sgs | |
| 172 | (Library.split_list (map (apsnd export_solution o fix_alls ctxt 1) subgoals)) | |
| 173 | end; | |
| 23171 | 174 | |
| 175 | end; |