src/HOL/Tools/Function/termination.ML
changeset 32705 04ce6bb14d85
parent 32683 7c1fe854ca6a
child 33099 b8cdd3d73022
equal deleted inserted replaced
32682:304a47739407 32705:04ce6bb14d85
   143           sk
   143           sk
   144   |> fst
   144   |> fst
   145 
   145 
   146 fun mk_sum_skel rel =
   146 fun mk_sum_skel rel =
   147   let
   147   let
   148     val cs = FundefLib.dest_binop_list @{const_name Set.union} rel
   148     val cs = FundefLib.dest_binop_list @{const_name Lattices.sup} rel
   149     fun collect_pats (Const (@{const_name Collect}, _) $ Abs (_, _, c)) =
   149     fun collect_pats (Const (@{const_name Collect}, _) $ Abs (_, _, c)) =
   150       let
   150       let
   151         val (Const ("op &", _) $ (Const ("op =", _) $ _ $ (Const ("Pair", _) $ r $ l)) $ Gam)
   151         val (Const ("op &", _) $ (Const ("op =", _) $ _ $ (Const ("Pair", _) $ r $ l)) $ Gam)
   152           = Term.strip_qnt_body "Ex" c
   152           = Term.strip_qnt_body "Ex" c
   153       in cons r o cons l end
   153       in cons r o cons l end
   231   end
   231   end
   232 
   232 
   233 fun CALLS tac i st =
   233 fun CALLS tac i st =
   234   if Thm.no_prems st then all_tac st
   234   if Thm.no_prems st then all_tac st
   235   else case Thm.term_of (Thm.cprem_of st i) of
   235   else case Thm.term_of (Thm.cprem_of st i) of
   236     (_ $ (_ $ rel)) => tac (FundefLib.dest_binop_list @{const_name Set.union} rel, i) st
   236     (_ $ (_ $ rel)) => tac (FundefLib.dest_binop_list @{const_name Lattices.sup} rel, i) st
   237   |_ => no_tac st
   237   |_ => no_tac st
   238 
   238 
   239 type ttac = (data -> int -> tactic) -> (data -> int -> tactic) -> data -> int -> tactic
   239 type ttac = (data -> int -> tactic) -> (data -> int -> tactic) -> data -> int -> tactic
   240 
   240 
   241 fun TERMINATION ctxt tac =
   241 fun TERMINATION ctxt tac =
   291 
   291 
   292       val relation =
   292       val relation =
   293           if null ineqs then
   293           if null ineqs then
   294               Const (@{const_name Set.empty}, fastype_of rel)
   294               Const (@{const_name Set.empty}, fastype_of rel)
   295           else
   295           else
   296               foldr1 (HOLogic.mk_binop @{const_name Set.union}) (map mk_compr ineqs)
   296               foldr1 (HOLogic.mk_binop @{const_name Lattices.sup}) (map mk_compr ineqs)
   297 
   297 
   298       fun solve_membership_tac i =
   298       fun solve_membership_tac i =
   299           (EVERY' (replicate (i - 2) (rtac @{thm UnI2}))  (* pick the right component of the union *)
   299           (EVERY' (replicate (i - 2) (rtac @{thm UnI2}))  (* pick the right component of the union *)
   300           THEN' (fn j => TRY (rtac @{thm UnI1} j))
   300           THEN' (fn j => TRY (rtac @{thm UnI1} j))
   301           THEN' (rtac @{thm CollectI})                    (* unfold comprehension *)
   301           THEN' (rtac @{thm CollectI})                    (* unfold comprehension *)