src/HOL/Tools/inductive_set.ML
changeset 45979 296d9a9c8d24
parent 45384 dffa657f0aa2
child 46219 426ed18eba43
equal deleted inserted replaced
45978:d3325de5f299 45979:296d9a9c8d24
   242     val thm' = Thm.instantiate ([], insts) thm;
   242     val thm' = Thm.instantiate ([], insts) thm;
   243     val thm'' = (case optfs' of
   243     val thm'' = (case optfs' of
   244         NONE => thm' RS sym
   244         NONE => thm' RS sym
   245       | SOME fs' =>
   245       | SOME fs' =>
   246           let
   246           let
   247             val U = List.last (binder_types T);
   247             val U = HOLogic.dest_setT (body_type T);
   248             val Ts = HOLogic.strip_ptupleT fs' U;
   248             val Ts = HOLogic.strip_ptupleT fs' U;
   249             (* FIXME: should cterm_instantiate increment indexes? *)
   249             (* FIXME: should cterm_instantiate increment indexes? *)
   250             val arg_cong' = Thm.incr_indexes (Thm.maxidx_of thm + 1) arg_cong;
   250             val arg_cong' = Thm.incr_indexes (Thm.maxidx_of thm + 1) arg_cong;
   251             val (arg_cong_f, _) = arg_cong' |> cprop_of |> Drule.strip_imp_concl |>
   251             val (arg_cong_f, _) = arg_cong' |> cprop_of |> Drule.strip_imp_concl |>
   252               Thm.dest_comb |> snd |> Drule.strip_comb |> snd |> hd |> Thm.dest_comb
   252               Thm.dest_comb |> snd |> Drule.strip_comb |> snd |> hd |> Thm.dest_comb
   349     eta_contract_thm (is_pred pred_arities) |>
   349     eta_contract_thm (is_pred pred_arities) |>
   350     Rule_Cases.save thm
   350     Rule_Cases.save thm
   351   end;
   351   end;
   352 
   352 
   353 val to_pred_att = Thm.rule_attribute o to_pred;
   353 val to_pred_att = Thm.rule_attribute o to_pred;
   354     
   354 
   355 
   355 
   356 (**** convert theorem in predicate notation to set notation ****)
   356 (**** convert theorem in predicate notation to set notation ****)
   357 
   357 
   358 fun to_set thms ctxt thm =
   358 fun to_set thms ctxt thm =
   359   let
   359   let
   419       | infer (Const (@{const_name Set.member}, _) $ t $ u) =
   419       | infer (Const (@{const_name Set.member}, _) $ t $ u) =
   420           infer_arities thy set_arities (SOME (HOLogic.flat_tuple_paths t), u)
   420           infer_arities thy set_arities (SOME (HOLogic.flat_tuple_paths t), u)
   421       | infer (t $ u) = infer t #> infer u
   421       | infer (t $ u) = infer t #> infer u
   422       | infer _ = I;
   422       | infer _ = I;
   423     val new_arities = filter_out
   423     val new_arities = filter_out
   424       (fn (x as Free (_, T), _) => member (op =) params x andalso length (binder_types T) > 1
   424       (fn (x as Free (_, T), _) => member (op =) params x andalso length (binder_types T) > 0
   425         | _ => false) (fold (snd #> infer) intros []);
   425         | _ => false) (fold (snd #> infer) intros []);
   426     val params' = map (fn x =>
   426     val params' = map (fn x =>
   427       (case AList.lookup op = new_arities x of
   427       (case AList.lookup op = new_arities x of
   428         SOME fs =>
   428         SOME fs =>
   429           let
   429           let
   445 
   445 
   446     (* equations for converting sets to predicates *)
   446     (* equations for converting sets to predicates *)
   447     val ((cs', cs_info), eqns) = cs |> map (fn c as Free (s, T) =>
   447     val ((cs', cs_info), eqns) = cs |> map (fn c as Free (s, T) =>
   448       let
   448       let
   449         val fs = the_default [] (AList.lookup op = new_arities c);
   449         val fs = the_default [] (AList.lookup op = new_arities c);
   450         val (Us, U) = split_last (binder_types T);
   450         val (Us, U) = strip_type T |> apsnd HOLogic.dest_setT;
   451         val _ = Us = paramTs orelse error (Pretty.string_of (Pretty.chunks
   451         val _ = Us = paramTs orelse error (Pretty.string_of (Pretty.chunks
   452           [Pretty.str "Argument types",
   452           [Pretty.str "Argument types",
   453            Pretty.block (Pretty.commas (map (Syntax.pretty_typ lthy) Us)),
   453            Pretty.block (Pretty.commas (map (Syntax.pretty_typ lthy) Us)),
   454            Pretty.str ("of " ^ s ^ " do not agree with types"),
   454            Pretty.str ("of " ^ s ^ " do not agree with types"),
   455            Pretty.block (Pretty.commas (map (Syntax.pretty_typ lthy) paramTs)),
   455            Pretty.block (Pretty.commas (map (Syntax.pretty_typ lthy) paramTs)),