src/HOL/Tools/Function/termination.ML
changeset 38864 4abe644fcea5
parent 38795 848be46708dc
child 39923 0e1bd289c8ea
equal deleted inserted replaced
38859:053c69cb4a0e 38864:4abe644fcea5
   146 fun mk_sum_skel rel =
   146 fun mk_sum_skel rel =
   147   let
   147   let
   148     val cs = Function_Lib.dest_binop_list @{const_name Lattices.sup} rel
   148     val cs = Function_Lib.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 (@{const_name HOL.conj}, _) $ (Const (@{const_name "op ="}, _) $ _ $ (Const (@{const_name Pair}, _) $ r $ l)) $ _)
   151         val (Const (@{const_name HOL.conj}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ (Const (@{const_name Pair}, _) $ r $ l)) $ _)
   152           = Term.strip_qnt_body @{const_name Ex} c
   152           = Term.strip_qnt_body @{const_name Ex} c
   153       in cons r o cons l end
   153       in cons r o cons l end
   154   in
   154   in
   155     mk_skel (fold collect_pats cs [])
   155     mk_skel (fold collect_pats cs [])
   156   end
   156   end
   183   let
   183   let
   184     val (sk, _, _, _, _) = D
   184     val (sk, _, _, _, _) = D
   185     val vs = Term.strip_qnt_vars @{const_name Ex} c
   185     val vs = Term.strip_qnt_vars @{const_name Ex} c
   186 
   186 
   187     (* FIXME: throw error "dest_call" for malformed terms *)
   187     (* FIXME: throw error "dest_call" for malformed terms *)
   188     val (Const (@{const_name HOL.conj}, _) $ (Const (@{const_name "op ="}, _) $ _ $ (Const (@{const_name Pair}, _) $ r $ l)) $ Gam)
   188     val (Const (@{const_name HOL.conj}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ (Const (@{const_name Pair}, _) $ r $ l)) $ Gam)
   189       = Term.strip_qnt_body @{const_name Ex} c
   189       = Term.strip_qnt_body @{const_name Ex} c
   190     val (p, l') = dest_inj sk l
   190     val (p, l') = dest_inj sk l
   191     val (q, r') = dest_inj sk r
   191     val (q, r') = dest_inj sk r
   192   in
   192   in
   193     (vs, p, l', q, r', Gam)
   193     (vs, p, l', q, r', Gam)