| author | haftmann | 
| Wed, 03 Mar 2010 17:21:45 +0100 | |
| changeset 35550 | e2bc7f8d8d51 | 
| parent 35408 | b48ab741683b | 
| child 36943 | ae740b96b914 | 
| permissions | -rw-r--r-- | 
| 29269 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
29064diff
changeset | 1 | (* Title: HOL/Statespace/distinct_tree_prover.ML | 
| 25171 | 2 | Author: Norbert Schirmer, TU Muenchen | 
| 3 | *) | |
| 4 | ||
| 25408 | 5 | signature DISTINCT_TREE_PROVER = | 
| 6 | sig | |
| 7 | datatype Direction = Left | Right | |
| 29302 | 8 |   val mk_tree : ('a -> term) -> typ -> 'a list -> term
 | 
| 9 | val dest_tree : term -> term list | |
| 10 | val find_tree : term -> term -> Direction list option | |
| 25408 | 11 | |
| 29302 | 12 | val neq_to_eq_False : thm | 
| 13 | val distinctTreeProver : thm -> Direction list -> Direction list -> thm | |
| 14 | val neq_x_y : Proof.context -> term -> term -> string -> thm option | |
| 15 | val distinctFieldSolver : string list -> solver | |
| 16 | val distinctTree_tac : string list -> Proof.context -> term * int -> tactic | |
| 17 | val distinct_implProver : thm -> cterm -> thm | |
| 18 | val subtractProver : term -> cterm -> thm -> thm | |
| 19 | val distinct_simproc : string list -> simproc | |
| 25408 | 20 | |
| 29302 | 21 | val discharge : thm list -> thm -> thm | 
| 25408 | 22 | end; | 
| 23 | ||
| 24 | structure DistinctTreeProver : DISTINCT_TREE_PROVER = | |
| 25171 | 25 | struct | 
| 26 | val all_distinct_left = thm "DistinctTreeProver.all_distinct_left"; | |
| 27 | val all_distinct_right = thm "DistinctTreeProver.all_distinct_right"; | |
| 28 | ||
| 29 | val distinct_left = thm "DistinctTreeProver.distinct_left"; | |
| 30 | val distinct_right = thm "DistinctTreeProver.distinct_right"; | |
| 31 | val distinct_left_right = thm "DistinctTreeProver.distinct_left_right"; | |
| 32 | val in_set_root = thm "DistinctTreeProver.in_set_root"; | |
| 33 | val in_set_left = thm "DistinctTreeProver.in_set_left"; | |
| 34 | val in_set_right = thm "DistinctTreeProver.in_set_right"; | |
| 35 | ||
| 36 | val swap_neq = thm "DistinctTreeProver.swap_neq"; | |
| 37 | val neq_to_eq_False = thm "DistinctTreeProver.neq_to_eq_False" | |
| 38 | ||
| 25408 | 39 | datatype Direction = Left | Right | 
| 40 | ||
| 25171 | 41 | fun treeT T = Type ("DistinctTreeProver.tree",[T]);
 | 
| 42 | fun mk_tree' e T n []     = Const ("DistinctTreeProver.tree.Tip",treeT T)
 | |
| 43 | | mk_tree' e T n xs = | |
| 44 | let | |
| 45 | val m = (n - 1) div 2; | |
| 46 | val (xsl,x::xsr) = chop m xs; | |
| 47 | val l = mk_tree' e T m xsl; | |
| 48 | val r = mk_tree' e T (n-(m+1)) xsr; | |
| 49 |      in Const ("DistinctTreeProver.tree.Node",
 | |
| 50 | treeT T --> T --> HOLogic.boolT--> treeT T --> treeT T) $ | |
| 51 | l$ e x $ HOLogic.false_const $ r | |
| 52 | end | |
| 53 | ||
| 54 | fun mk_tree e T xs = mk_tree' e T (length xs) xs; | |
| 55 | ||
| 56 | fun dest_tree (Const ("DistinctTreeProver.tree.Tip",_)) = []
 | |
| 57 |   | dest_tree (Const ("DistinctTreeProver.tree.Node",_)$l$e$_$r) = dest_tree l @ e :: dest_tree r
 | |
| 58 |   | dest_tree t = raise TERM ("DistinctTreeProver.dest_tree",[t]);
 | |
| 59 | ||
| 25408 | 60 | |
| 25171 | 61 | |
| 62 | fun lin_find_tree e (Const ("DistinctTreeProver.tree.Tip",_)) = NONE
 | |
| 63 |   | lin_find_tree e (Const ("DistinctTreeProver.tree.Node",_) $ l$ x $ _ $ r) =
 | |
| 64 | if e aconv x | |
| 65 | then SOME [] | |
| 66 | else (case lin_find_tree e l of | |
| 67 | SOME path => SOME (Left::path) | |
| 68 | | NONE => (case lin_find_tree e r of | |
| 69 | SOME path => SOME (Right::path) | |
| 70 | | NONE => NONE)) | |
| 71 |   | lin_find_tree e t = raise TERM ("find_tree: input not a tree",[t])
 | |
| 72 | ||
| 73 | fun bin_find_tree order e (Const ("DistinctTreeProver.tree.Tip",_)) = NONE
 | |
| 74 |   | bin_find_tree order e (Const ("DistinctTreeProver.tree.Node",_) $ l$ x $ _ $ r) =
 | |
| 75 | (case order (e,x) of | |
| 76 | EQUAL => SOME [] | |
| 77 | | LESS => Option.map (cons Left) (bin_find_tree order e l) | |
| 78 | | GREATER => Option.map (cons Right) (bin_find_tree order e r)) | |
| 79 |   | bin_find_tree order e t = raise TERM ("find_tree: input not a tree",[t])
 | |
| 80 | ||
| 81 | fun find_tree e t = | |
| 35408 | 82 | (case bin_find_tree Term_Ord.fast_term_ord e t of | 
| 25171 | 83 | NONE => lin_find_tree e t | 
| 84 | | x => x); | |
| 85 | ||
| 86 | ||
| 87 | fun index_tree (Const ("DistinctTreeProver.tree.Tip",_)) path tab = tab
 | |
| 88 |   | index_tree (Const ("DistinctTreeProver.tree.Node",_) $ l$ x $ _ $ r) path tab =
 | |
| 89 | tab | |
| 90 | |> Termtab.update_new (x,path) | |
| 91 | |> index_tree l (path@[Left]) | |
| 92 | |> index_tree r (path@[Right]) | |
| 93 |   | index_tree t _ _ = raise TERM ("index_tree: input not a tree",[t])
 | |
| 94 | ||
| 95 | fun split_common_prefix xs [] = ([],xs,[]) | |
| 96 | | split_common_prefix [] ys = ([],[],ys) | |
| 97 | | split_common_prefix (xs as (x::xs')) (ys as (y::ys')) = | |
| 98 | if x=y | |
| 99 | then let val (ps,xs'',ys'') = split_common_prefix xs' ys' in (x::ps,xs'',ys'') end | |
| 100 | else ([],xs,ys) | |
| 101 | ||
| 102 | ||
| 103 | (* Wrapper around Thm.instantiate. The type instiations of instTs are applied to | |
| 104 | * the right hand sides of insts | |
| 105 | *) | |
| 106 | fun instantiate instTs insts = | |
| 107 | let | |
| 108 | val instTs' = map (fn (T,U) => (dest_TVar (typ_of T),typ_of U)) instTs; | |
| 109 | fun substT x = (case AList.lookup (op =) instTs' x of NONE => TVar x | SOME T' => T'); | |
| 110 | fun mapT_and_recertify ct = | |
| 111 | let | |
| 112 | val thy = theory_of_cterm ct; | |
| 113 | in (cterm_of thy (Term.map_types (Term.map_type_tvar substT) (term_of ct))) end; | |
| 114 | val insts' = map (apfst mapT_and_recertify) insts; | |
| 115 | in Thm.instantiate (instTs,insts') end; | |
| 116 | ||
| 117 | fun tvar_clash ixn S S' = raise TYPE ("Type variable " ^
 | |
| 118 | quote (Term.string_of_vname ixn) ^ " has two distinct sorts", | |
| 119 | [TVar (ixn, S), TVar (ixn, S')], []); | |
| 120 | ||
| 121 | fun lookup (tye, (ixn, S)) = | |
| 122 | (case AList.lookup (op =) tye ixn of | |
| 123 | NONE => NONE | |
| 124 | | SOME (S', T) => if S = S' then SOME T else tvar_clash ixn S S'); | |
| 125 | ||
| 126 | val naive_typ_match = | |
| 127 | let | |
| 128 | fun match (TVar (v, S), T) subs = | |
| 129 | (case lookup (subs, (v, S)) of | |
| 130 | NONE => ((v, (S, T))::subs) | |
| 131 | | SOME _ => subs) | |
| 132 | | match (Type (a, Ts), Type (b, Us)) subs = | |
| 133 | if a <> b then raise Type.TYPE_MATCH | |
| 134 | else matches (Ts, Us) subs | |
| 135 | | match (TFree x, TFree y) subs = | |
| 136 | if x = y then subs else raise Type.TYPE_MATCH | |
| 137 | | match _ _ = raise Type.TYPE_MATCH | |
| 138 | and matches (T :: Ts, U :: Us) subs = matches (Ts, Us) (match (T, U) subs) | |
| 139 | | matches _ subs = subs; | |
| 140 | in match end; | |
| 141 | ||
| 142 | ||
| 143 | (* expects that relevant type variables are already contained in | |
| 144 | * term variables. First instantiation of variables is returned without further | |
| 145 | * checking. | |
| 146 | *) | |
| 147 | fun naive_cterm_first_order_match (t,ct) env = | |
| 148 | let | |
| 149 | val thy = (theory_of_cterm ct); | |
| 150 | fun mtch (env as (tyinsts,insts)) = fn | |
| 151 | (Var(ixn,T),ct) => | |
| 152 | (case AList.lookup (op =) insts ixn of | |
| 153 | NONE => (naive_typ_match (T,typ_of (ctyp_of_term ct)) tyinsts, | |
| 154 | (ixn, ct)::insts) | |
| 155 | | SOME _ => env) | |
| 156 | | (f$t,ct) => let val (cf,ct') = Thm.dest_comb ct; | |
| 157 | in mtch (mtch env (f,cf)) (t,ct') end | |
| 158 | | _ => env | |
| 159 | in mtch env (t,ct) end; | |
| 160 | ||
| 161 | ||
| 162 | fun mp prem rule = implies_elim rule prem; | |
| 163 | ||
| 164 | fun discharge prems rule = | |
| 165 | let | |
| 166 | val thy = theory_of_thm (hd prems); | |
| 167 | val (tyinsts,insts) = | |
| 168 | fold naive_cterm_first_order_match (prems_of rule ~~ map cprop_of prems) ([],[]); | |
| 169 | ||
| 170 | val tyinsts' = map (fn (v,(S,U)) => (ctyp_of thy (TVar (v,S)),ctyp_of thy U)) | |
| 171 | tyinsts; | |
| 172 | val insts' = map (fn (idxn,ct) => (cterm_of thy (Var (idxn,typ_of (ctyp_of_term ct))),ct)) | |
| 173 | insts; | |
| 174 | val rule' = Thm.instantiate (tyinsts',insts') rule; | |
| 175 | in fold mp prems rule' end; | |
| 176 | ||
| 177 | local | |
| 178 | ||
| 179 | val (l_in_set_root,x_in_set_root,r_in_set_root) = | |
| 180 | let val (Node_l_x_d,r) = (cprop_of in_set_root) | |
| 181 | |> Thm.dest_comb |> #2 | |
| 182 | |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2 |> Thm.dest_comb; | |
| 183 | val (Node_l,x) = Node_l_x_d |> Thm.dest_comb |> #1 |> Thm.dest_comb; | |
| 184 | val l = Node_l |> Thm.dest_comb |> #2; | |
| 185 | in (l,x,r) end | |
| 186 | val (x_in_set_left,r_in_set_left) = | |
| 187 | let val (Node_l_x_d,r) = (cprop_of in_set_left) | |
| 188 | |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2 | |
| 189 | |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2 |> Thm.dest_comb; | |
| 190 | val x = Node_l_x_d |> Thm.dest_comb |> #1 |> Thm.dest_comb |> #2; | |
| 191 | in (x,r) end | |
| 192 | ||
| 193 | val (x_in_set_right,l_in_set_right) = | |
| 194 | let val (Node_l,x) = (cprop_of in_set_right) | |
| 195 | |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2 | |
| 196 | |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2 | |
| 197 | |> Thm.dest_comb |> #1 |> Thm.dest_comb |> #1 | |
| 198 | |> Thm.dest_comb | |
| 199 | val l = Node_l |> Thm.dest_comb |> #2; | |
| 200 | in (x,l) end | |
| 201 | ||
| 202 | in | |
| 203 | (* | |
| 204 | 1. First get paths x_path y_path of x and y in the tree. | |
| 205 | 2. For the common prefix descend into the tree according to the path | |
| 206 | and lemmas all_distinct_left/right | |
| 207 | 3. If one restpath is empty use distinct_left/right, | |
| 208 | otherwise all_distinct_left_right | |
| 209 | *) | |
| 210 | ||
| 211 | fun distinctTreeProver dist_thm x_path y_path = | |
| 212 | let | |
| 213 | fun dist_subtree [] thm = thm | |
| 214 | | dist_subtree (p::ps) thm = | |
| 215 | let | |
| 216 | val rule = (case p of Left => all_distinct_left | Right => all_distinct_right) | |
| 217 | in dist_subtree ps (discharge [thm] rule) end; | |
| 218 | ||
| 219 | val (ps,x_rest,y_rest) = split_common_prefix x_path y_path; | |
| 220 | val dist_subtree_thm = dist_subtree ps dist_thm; | |
| 221 | val subtree = cprop_of dist_subtree_thm |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2; | |
| 222 | val (_,[l,_,_,r]) = Drule.strip_comb subtree; | |
| 223 | ||
| 224 | fun in_set ps tree = | |
| 225 | let | |
| 226 | val (_,[l,x,_,r]) = Drule.strip_comb tree; | |
| 227 | val xT = ctyp_of_term x; | |
| 228 | in (case ps of | |
| 229 | [] => instantiate | |
| 230 | [(ctyp_of_term x_in_set_root,xT)] | |
| 231 | [(l_in_set_root,l),(x_in_set_root,x),(r_in_set_root,r)] in_set_root | |
| 232 | | (Left::ps') => | |
| 233 | let | |
| 234 | val in_set_l = in_set ps' l; | |
| 235 | val in_set_left' = instantiate [(ctyp_of_term x_in_set_left,xT)] | |
| 236 | [(x_in_set_left,x),(r_in_set_left,r)] in_set_left; | |
| 237 | in discharge [in_set_l] in_set_left' end | |
| 238 | | (Right::ps') => | |
| 239 | let | |
| 240 | val in_set_r = in_set ps' r; | |
| 241 | val in_set_right' = instantiate [(ctyp_of_term x_in_set_right,xT)] | |
| 242 | [(x_in_set_right,x),(l_in_set_right,l)] in_set_right; | |
| 243 | in discharge [in_set_r] in_set_right' end) | |
| 244 | end | |
| 245 | ||
| 246 |   fun in_set' [] = raise TERM ("distinctTreeProver",[])
 | |
| 247 | | in_set' (Left::ps) = in_set ps l | |
| 248 | | in_set' (Right::ps) = in_set ps r; | |
| 249 | ||
| 250 | fun distinct_lr node_in_set Left = discharge [dist_subtree_thm,node_in_set] distinct_left | |
| 251 | | distinct_lr node_in_set Right = discharge [dist_subtree_thm,node_in_set] distinct_right | |
| 252 | ||
| 253 | val (swap,neq) = | |
| 254 | (case x_rest of | |
| 255 | [] => let | |
| 256 | val y_in_set = in_set' y_rest; | |
| 257 | in (false,distinct_lr y_in_set (hd y_rest)) end | |
| 258 | | (xr::xrs) => | |
| 259 | (case y_rest of | |
| 260 | [] => let | |
| 261 | val x_in_set = in_set' x_rest; | |
| 262 | in (true,distinct_lr x_in_set (hd x_rest)) end | |
| 263 | | (yr::yrs) => | |
| 264 | let | |
| 265 | val x_in_set = in_set' x_rest; | |
| 266 | val y_in_set = in_set' y_rest; | |
| 267 | in (case xr of | |
| 268 | Left => (false, | |
| 269 | discharge [dist_subtree_thm,x_in_set,y_in_set] distinct_left_right) | |
| 270 | |Right => (true, | |
| 271 | discharge [dist_subtree_thm,y_in_set,x_in_set] distinct_left_right)) | |
| 272 | end | |
| 273 | )) | |
| 274 | in if swap then discharge [neq] swap_neq else neq | |
| 275 | end | |
| 276 | ||
| 277 | ||
| 278 | val delete_root = thm "DistinctTreeProver.delete_root"; | |
| 279 | val delete_left = thm "DistinctTreeProver.delete_left"; | |
| 280 | val delete_right = thm "DistinctTreeProver.delete_right"; | |
| 281 | ||
| 282 | fun deleteProver dist_thm [] = delete_root OF [dist_thm] | |
| 283 | | deleteProver dist_thm (p::ps) = | |
| 284 | let | |
| 285 | val dist_rule = (case p of Left => all_distinct_left | Right => all_distinct_right); | |
| 286 | val dist_thm' = discharge [dist_thm] dist_rule | |
| 287 | val del_rule = (case p of Left => delete_left | Right => delete_right) | |
| 288 | val del = deleteProver dist_thm' ps; | |
| 289 | in discharge [dist_thm, del] del_rule end; | |
| 290 | ||
| 291 | val subtract_Tip = thm "DistinctTreeProver.subtract_Tip"; | |
| 292 | val subtract_Node = thm "DistinctTreeProver.subtract_Node"; | |
| 293 | val delete_Some_all_distinct = thm "DistinctTreeProver.delete_Some_all_distinct"; | |
| 294 | val subtract_Some_all_distinct_res = thm "DistinctTreeProver.subtract_Some_all_distinct_res"; | |
| 295 | ||
| 296 | local | |
| 297 | val (alpha,v) = | |
| 298 | let | |
| 299 | val ct = subtract_Tip |> Thm.cprop_of |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2 | |
| 300 | |> Thm.dest_comb |> #2 | |
| 301 | val [alpha] = ct |> Thm.ctyp_of_term |> Thm.dest_ctyp; | |
| 302 | in (alpha, #1 (dest_Var (term_of ct))) end; | |
| 303 | in | |
| 304 | fun subtractProver (Const ("DistinctTreeProver.tree.Tip",T)) ct dist_thm =
 | |
| 305 | let | |
| 306 | val ct' = dist_thm |> Thm.cprop_of |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2; | |
| 307 | val thy = theory_of_cterm ct; | |
| 308 | val [alphaI] = #2 (dest_Type T); | |
| 309 | in Thm.instantiate ([(alpha,ctyp_of thy alphaI)], | |
| 310 | [(cterm_of thy (Var (v,treeT alphaI)),ct')]) subtract_Tip | |
| 311 | end | |
| 312 |   | subtractProver (Const ("DistinctTreeProver.tree.Node",nT)$l$x$d$r) ct dist_thm =
 | |
| 313 | let | |
| 314 | val ct' = dist_thm |> Thm.cprop_of |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2; | |
| 315 | val (_,[cl,_,_,cr]) = Drule.strip_comb ct; | |
| 316 | val ps = the (find_tree x (term_of ct')); | |
| 317 | val del_tree = deleteProver dist_thm ps; | |
| 318 | val dist_thm' = discharge [del_tree, dist_thm] delete_Some_all_distinct; | |
| 319 | val sub_l = subtractProver (term_of cl) cl (dist_thm'); | |
| 320 | val sub_r = subtractProver (term_of cr) cr | |
| 321 | (discharge [sub_l, dist_thm'] subtract_Some_all_distinct_res); | |
| 322 | in discharge [del_tree, sub_l, sub_r] subtract_Node end | |
| 323 | end | |
| 324 | ||
| 325 | val subtract_Some_all_distinct = thm "DistinctTreeProver.subtract_Some_all_distinct"; | |
| 326 | fun distinct_implProver dist_thm ct = | |
| 327 | let | |
| 328 | val ctree = ct |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2; | |
| 329 | val sub = subtractProver (term_of ctree) ctree dist_thm; | |
| 330 | in subtract_Some_all_distinct OF [sub, dist_thm] end; | |
| 331 | ||
| 332 | fun get_fst_success f [] = NONE | |
| 333 | | get_fst_success f (x::xs) = (case f x of NONE => get_fst_success f xs | |
| 334 | | SOME v => SOME v); | |
| 335 | ||
| 336 | fun neq_x_y ctxt x y name = | |
| 337 | (let | |
| 26343 
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
 wenzelm parents: 
26336diff
changeset | 338 | val dist_thm = the (try (ProofContext.get_thm ctxt) name); | 
| 25171 | 339 | val ctree = cprop_of dist_thm |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2; | 
| 340 | val tree = term_of ctree; | |
| 341 | val x_path = the (find_tree x tree); | |
| 342 | val y_path = the (find_tree y tree); | |
| 343 | val thm = distinctTreeProver dist_thm x_path y_path; | |
| 344 | in SOME thm | |
| 345 | end handle Option => NONE) | |
| 346 | ||
| 347 | fun distinctTree_tac names ctxt | |
| 348 |       (Const ("Trueprop",_) $ (Const ("Not", _) $ (Const ("op =", _) $ x $ y)), i) =
 | |
| 349 | (case get_fst_success (neq_x_y ctxt x y) names of | |
| 350 | SOME neq => rtac neq i | |
| 351 | | NONE => no_tac) | |
| 352 | | distinctTree_tac _ _ _ = no_tac; | |
| 353 | ||
| 354 | fun distinctFieldSolver names = mk_solver' "distinctFieldSolver" | |
| 30289 | 355 | (fn ss => case try Simplifier.the_context ss of | 
| 25171 | 356 | SOME ctxt => SUBGOAL (distinctTree_tac names ctxt) | 
| 357 | | NONE => fn i => no_tac) | |
| 358 | ||
| 359 | fun distinct_simproc names = | |
| 29064 | 360 |   Simplifier.simproc @{theory HOL} "DistinctTreeProver.distinct_simproc" ["x = y"]
 | 
| 25171 | 361 |     (fn thy => fn ss => fn (Const ("op =",_)$x$y) =>
 | 
| 30289 | 362 | case try Simplifier.the_context ss of | 
| 25171 | 363 | SOME ctxt => Option.map (fn neq => neq_to_eq_False OF [neq]) | 
| 364 | (get_fst_success (neq_x_y ctxt x y) names) | |
| 365 | | NONE => NONE | |
| 366 | ) | |
| 367 | ||
| 368 | end | |
| 369 | end; |