| author | immler | 
| Tue, 27 Aug 2019 22:41:47 +0200 | |
| changeset 70617 | c81ac117afa6 | 
| parent 67522 | 9e712280cc37 | 
| child 74282 | c2ee8d993d6a | 
| permissions | -rw-r--r-- | 
| 30160 
5f7b17941730
moved some generic tools to src/Tools/ -- src/Provers is essentially obsolete;
 wenzelm parents: 
29273diff
changeset | 1 | (* Title: Tools/coherent.ML | 
| 28326 | 2 | Author: Stefan Berghofer, TU Muenchen | 
| 31241 | 3 | Author: Marc Bezem, Institutt for Informatikk, Universitetet i Bergen | 
| 28326 | 4 | |
| 5 | Prover for coherent logic, see e.g. | |
| 6 | ||
| 7 | Marc Bezem and Thierry Coquand, Automating Coherent Logic, LPAR 2005 | |
| 8 | ||
| 9 | for a description of the algorithm. | |
| 10 | *) | |
| 11 | ||
| 12 | signature COHERENT_DATA = | |
| 13 | sig | |
| 14 | val atomize_elimL: thm | |
| 15 | val atomize_exL: thm | |
| 16 | val atomize_conjL: thm | |
| 17 | val atomize_disjL: thm | |
| 18 | val operator_names: string list | |
| 19 | end; | |
| 20 | ||
| 21 | signature COHERENT = | |
| 22 | sig | |
| 55632 | 23 | val trace: bool Config.T | 
| 31241 | 24 | val coherent_tac: Proof.context -> thm list -> int -> tactic | 
| 28326 | 25 | end; | 
| 26 | ||
| 32734 | 27 | functor Coherent(Data: COHERENT_DATA) : COHERENT = | 
| 28326 | 28 | struct | 
| 29 | ||
| 31241 | 30 | (** misc tools **) | 
| 31 | ||
| 67149 | 32 | val (trace, trace_setup) = Attrib.config_bool \<^binding>\<open>coherent_trace\<close> (K false); | 
| 55632 | 33 | fun cond_trace ctxt msg = if Config.get ctxt trace then tracing (msg ()) else (); | 
| 28326 | 34 | |
| 35 | datatype cl_prf = | |
| 36 | ClPrf of thm * (Type.tyenv * Envir.tenv) * ((indexname * typ) * term) list * | |
| 37 | int list * (term list * cl_prf) list; | |
| 38 | ||
| 29273 | 39 | val is_atomic = not o exists_Const (member (op =) Data.operator_names o #1); | 
| 28326 | 40 | |
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
46186diff
changeset | 41 | fun rulify_elim_conv ctxt ct = | 
| 59582 | 42 | if is_atomic (Logic.strip_imp_concl (Thm.term_of ct)) then Conv.all_conv ct | 
| 43 | else Conv.concl_conv (length (Logic.strip_imp_prems (Thm.term_of ct))) | |
| 36946 | 44 | (Conv.rewr_conv (Thm.symmetric Data.atomize_elimL) then_conv | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
46186diff
changeset | 45 | Raw_Simplifier.rewrite ctxt true (map Thm.symmetric | 
| 28326 | 46 | [Data.atomize_exL, Data.atomize_conjL, Data.atomize_disjL])) ct | 
| 47 | ||
| 54883 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 wenzelm parents: 
54742diff
changeset | 48 | fun rulify_elim ctxt th = Simplifier.norm_hhf ctxt (Conv.fconv_rule (rulify_elim_conv ctxt) th); | 
| 28326 | 49 | |
| 50 | (* Decompose elimination rule of the form | |
| 51 | A1 ==> ... ==> Am ==> (!!xs1. Bs1 ==> P) ==> ... ==> (!!xsn. Bsn ==> P) ==> P | |
| 52 | *) | |
| 53 | fun dest_elim prop = | |
| 54 | let | |
| 55 | val prems = Logic.strip_imp_prems prop; | |
| 56 | val concl = Logic.strip_imp_concl prop; | |
| 67522 | 57 | val (prems1, prems2) = chop_suffix (fn t => Logic.strip_assums_concl t = concl) prems; | 
| 28326 | 58 | in | 
| 59 | (prems1, | |
| 60 | if null prems2 then [([], [concl])] | |
| 61 | else map (fn t => | |
| 62 | (map snd (Logic.strip_params t), Logic.strip_assums_hyp t)) prems2) | |
| 63 | end; | |
| 64 | ||
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
46186diff
changeset | 65 | fun mk_rule ctxt th = | 
| 28326 | 66 | let | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
46186diff
changeset | 67 | val th' = rulify_elim ctxt th; | 
| 59582 | 68 | val (prems, cases) = dest_elim (Thm.prop_of th') | 
| 28326 | 69 | in (th', prems, cases) end; | 
| 70 | ||
| 71 | fun mk_dom ts = fold (fn t => | |
| 72 | Typtab.map_default (fastype_of t, []) (fn us => us @ [t])) ts Typtab.empty; | |
| 73 | ||
| 74 | val empty_env = (Vartab.empty, Vartab.empty); | |
| 75 | ||
| 76 | (* Find matcher that makes conjunction valid in given state *) | |
| 55631 | 77 | fun valid_conj _ _ env [] = Seq.single (env, []) | 
| 28326 | 78 | | valid_conj ctxt facts env (t :: ts) = | 
| 79 | Seq.maps (fn (u, x) => Seq.map (apsnd (cons x)) | |
| 80 | (valid_conj ctxt facts | |
| 42361 | 81 | (Pattern.match (Proof_Context.theory_of ctxt) (t, u) env) ts | 
| 28326 | 82 | handle Pattern.MATCH => Seq.empty)) | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58839diff
changeset | 83 | (Seq.of_list (sort (int_ord o apply2 snd) (Net.unify_term facts t))); | 
| 28326 | 84 | |
| 85 | (* Instantiate variables that only occur free in conlusion *) | |
| 86 | fun inst_extra_vars ctxt dom cs = | |
| 87 | let | |
| 88 | val vs = fold Term.add_vars (maps snd cs) []; | |
| 89 | fun insts [] inst = Seq.single inst | |
| 55627 | 90 | | insts ((ixn, T) :: vs') inst = | 
| 91 | Seq.maps (fn t => insts vs' (((ixn, T), t) :: inst)) | |
| 92 | (Seq.of_list | |
| 93 | (case Typtab.lookup dom T of | |
| 94 | NONE => | |
| 95 |                   error ("Unknown domain: " ^
 | |
| 96 | Syntax.string_of_typ ctxt T ^ "\nfor term(s) " ^ | |
| 97 | commas (maps (map (Syntax.string_of_term ctxt) o snd) cs)) | |
| 98 | | SOME ts => ts)) | |
| 99 | in | |
| 100 | Seq.map (fn inst => | |
| 101 | (inst, map (apsnd (map (subst_Vars (map (apfst fst) inst)))) cs)) | |
| 102 | (insts vs []) | |
| 28326 | 103 | end; | 
| 104 | ||
| 105 | (* Check whether disjunction is valid in given state *) | |
| 55631 | 106 | fun is_valid_disj _ _ [] = false | 
| 28326 | 107 | | is_valid_disj ctxt facts ((Ts, ts) :: ds) = | 
| 55627 | 108 |       let val vs = map_index (fn (i, T) => Var (("x", i), T)) Ts in
 | 
| 109 | (case Seq.pull (valid_conj ctxt facts empty_env | |
| 110 | (map (fn t => subst_bounds (rev vs, t)) ts)) of | |
| 28326 | 111 | SOME _ => true | 
| 55627 | 112 | | NONE => is_valid_disj ctxt facts ds) | 
| 28326 | 113 | end; | 
| 114 | ||
| 55627 | 115 | fun string_of_facts ctxt s facts = | 
| 55634 | 116 | Pretty.string_of (Pretty.big_list s | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58839diff
changeset | 117 | (map (Syntax.pretty_term ctxt) (map fst (sort (int_ord o apply2 snd) (Net.content facts))))); | 
| 28326 | 118 | |
| 119 | fun valid ctxt rules goal dom facts nfacts nparams = | |
| 55627 | 120 | let | 
| 121 | val seq = | |
| 122 | Seq.of_list rules |> Seq.maps (fn (th, ps, cs) => | |
| 123 | valid_conj ctxt facts empty_env ps |> Seq.maps (fn (env as (tye, _), is) => | |
| 124 | let val cs' = map (fn (Ts, ts) => | |
| 125 | (map (Envir.subst_type tye) Ts, map (Envir.subst_term env) ts)) cs | |
| 126 | in | |
| 127 | inst_extra_vars ctxt dom cs' |> | |
| 128 | Seq.map_filter (fn (inst, cs'') => | |
| 129 | if is_valid_disj ctxt facts cs'' then NONE | |
| 130 | else SOME (th, env, inst, is, cs'')) | |
| 131 | end)); | |
| 28326 | 132 | in | 
| 55627 | 133 | (case Seq.pull seq of | 
| 55634 | 134 | NONE => | 
| 135 | (if Context_Position.is_visible ctxt then | |
| 136 | warning (string_of_facts ctxt "Countermodel found:" facts) | |
| 137 | else (); NONE) | |
| 28326 | 138 | | SOME ((th, env, inst, is, cs), _) => | 
| 139 | if cs = [([], [goal])] then SOME (ClPrf (th, env, inst, is, [])) | |
| 140 | else | |
| 141 | (case valid_cases ctxt rules goal dom facts nfacts nparams cs of | |
| 55627 | 142 | NONE => NONE | 
| 143 | | SOME prfs => SOME (ClPrf (th, env, inst, is, prfs)))) | |
| 28326 | 144 | end | 
| 145 | ||
| 55631 | 146 | and valid_cases _ _ _ _ _ _ _ [] = SOME [] | 
| 28326 | 147 | | valid_cases ctxt rules goal dom facts nfacts nparams ((Ts, ts) :: ds) = | 
| 148 | let | |
| 55634 | 149 | val _ = | 
| 150 | cond_trace ctxt (fn () => | |
| 151 | Pretty.string_of (Pretty.block | |
| 152 | (Pretty.str "case" :: Pretty.brk 1 :: | |
| 153 | Pretty.commas (map (Syntax.pretty_term ctxt) ts)))); | |
| 55636 | 154 | |
| 155 |         val ps = map_index (fn (i, T) => ("par" ^ string_of_int (nparams + i), T)) Ts;
 | |
| 156 | val (params, ctxt') = fold_map Variable.next_bound ps ctxt; | |
| 55627 | 157 | val ts' = map_index (fn (i, t) => (subst_bounds (rev params, t), nfacts + i)) ts; | 
| 158 | val dom' = | |
| 159 | fold (fn (T, p) => Typtab.map_default (T, []) (fn ps => ps @ [p])) (Ts ~~ params) dom; | |
| 160 | val facts' = fold (fn (t, i) => Net.insert_term op = (t, (t, i))) ts' facts; | |
| 28326 | 161 | in | 
| 55636 | 162 | (case valid ctxt' rules goal dom' facts' (nfacts + length ts) (nparams + length Ts) of | 
| 28326 | 163 | NONE => NONE | 
| 55627 | 164 | | SOME prf => | 
| 165 | (case valid_cases ctxt rules goal dom facts nfacts nparams ds of | |
| 166 | NONE => NONE | |
| 167 | | SOME prfs => SOME ((params, prf) :: prfs))) | |
| 28326 | 168 | end; | 
| 169 | ||
| 31241 | 170 | |
| 28326 | 171 | (** proof replaying **) | 
| 172 | ||
| 55632 | 173 | fun thm_of_cl_prf ctxt goal asms (ClPrf (th, (tye, env), insts, is, prfs)) = | 
| 28326 | 174 | let | 
| 55627 | 175 | val _ = | 
| 55632 | 176 | cond_trace ctxt (fn () => | 
| 61268 | 177 | Pretty.string_of (Pretty.big_list "asms:" (map (Thm.pretty_thm ctxt) asms))); | 
| 55627 | 178 | val th' = | 
| 179 | Drule.implies_elim_list | |
| 180 | (Thm.instantiate | |
| 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: 
59642diff
changeset | 181 | (map (fn (ixn, (S, T)) => ((ixn, S), Thm.ctyp_of ctxt T)) (Vartab.dest tye), | 
| 55627 | 182 | map (fn (ixn, (T, t)) => | 
| 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: 
59642diff
changeset | 183 | ((ixn, Envir.subst_type tye T), Thm.cterm_of ctxt t)) (Vartab.dest env) @ | 
| 
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: 
59642diff
changeset | 184 | map (fn (ixnT, t) => (ixnT, Thm.cterm_of ctxt t)) insts) th) | 
| 55627 | 185 | (map (nth asms) is); | 
| 59582 | 186 | val (_, cases) = dest_elim (Thm.prop_of th'); | 
| 28326 | 187 | in | 
| 55627 | 188 | (case (cases, prfs) of | 
| 28326 | 189 | ([([], [_])], []) => th' | 
| 55632 | 190 | | ([([], [_])], [([], prf)]) => thm_of_cl_prf ctxt goal (asms @ [th']) prf | 
| 55627 | 191 | | _ => | 
| 192 | Drule.implies_elim_list | |
| 193 | (Thm.instantiate (Thm.match | |
| 59582 | 194 | (Drule.strip_imp_concl (Thm.cprop_of th'), goal)) th') | 
| 55632 | 195 | (map (thm_of_case_prf ctxt goal asms) (prfs ~~ cases))) | 
| 28326 | 196 | end | 
| 197 | ||
| 55632 | 198 | and thm_of_case_prf ctxt goal asms ((params, prf), (_, asms')) = | 
| 28326 | 199 | let | 
| 59642 | 200 | val cparams = map (Thm.cterm_of ctxt) params; | 
| 201 | val asms'' = map (Thm.cterm_of ctxt o curry subst_bounds (rev params)) asms'; | |
| 55634 | 202 | val (prems'', ctxt') = fold_map Thm.assume_hyps asms'' ctxt; | 
| 28326 | 203 | in | 
| 55627 | 204 | Drule.forall_intr_list cparams | 
| 55634 | 205 | (Drule.implies_intr_list asms'' (thm_of_cl_prf ctxt' goal (asms @ prems'') prf)) | 
| 28326 | 206 | end; | 
| 207 | ||
| 208 | ||
| 209 | (** external interface **) | |
| 210 | ||
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
46186diff
changeset | 211 | fun coherent_tac ctxt rules = SUBPROOF (fn {prems, concl, params, context = ctxt', ...} =>
 | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59058diff
changeset | 212 | resolve_tac ctxt' [rulify_elim_conv ctxt' concl RS Drule.equal_elim_rule2] 1 THEN | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
46186diff
changeset | 213 |   SUBPROOF (fn {prems = prems', concl, context = ctxt'', ...} =>
 | 
| 55627 | 214 | let | 
| 215 | val xs = | |
| 59582 | 216 | map (Thm.term_of o #2) params @ | 
| 55627 | 217 | map (fn (_, s) => Free (s, the (Variable.default_type ctxt'' s))) | 
| 218 | (rev (Variable.dest_fixes ctxt'')) (* FIXME !? *) | |
| 28326 | 219 | in | 
| 55632 | 220 | (case | 
| 59582 | 221 | valid ctxt'' (map (mk_rule ctxt'') (prems' @ prems @ rules)) (Thm.term_of concl) | 
| 55632 | 222 | (mk_dom xs) Net.empty 0 0 of | 
| 55627 | 223 | NONE => no_tac | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59058diff
changeset | 224 | | SOME prf => resolve_tac ctxt'' [thm_of_cl_prf ctxt'' concl [] prf] 1) | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
46186diff
changeset | 225 | end) ctxt' 1) ctxt; | 
| 28326 | 226 | |
| 55632 | 227 | val _ = Theory.setup | 
| 228 | (trace_setup #> | |
| 67149 | 229 | Method.setup \<^binding>\<open>coherent\<close> | 
| 55632 | 230 | (Attrib.thms >> (fn rules => fn ctxt => | 
| 231 | METHOD (fn facts => HEADGOAL (coherent_tac ctxt (facts @ rules))))) | |
| 232 | "prove coherent formula"); | |
| 28326 | 233 | |
| 234 | end; |