src/HOL/Tools/Function/partial_function.ML
changeset 43083 df41a5762c3d
parent 43080 73a1d6a7ef1d
child 45008 8b74cfea913a
equal deleted inserted replaced
43082:8d0c44de9773 43083:df41a5762c3d
   158 fun uncurry_n arity = funpow (arity - 1) HOLogic.mk_split;
   158 fun uncurry_n arity = funpow (arity - 1) HOLogic.mk_split;
   159 
   159 
   160 val curry_uncurry_ss = HOL_basic_ss addsimps
   160 val curry_uncurry_ss = HOL_basic_ss addsimps
   161   [@{thm Product_Type.curry_split}, @{thm Product_Type.split_curry}]
   161   [@{thm Product_Type.curry_split}, @{thm Product_Type.split_curry}]
   162 
   162 
       
   163 val split_conv_ss = HOL_basic_ss addsimps
       
   164   [@{thm Product_Type.split_conv}];
       
   165 
       
   166 fun mk_curried_induct args ctxt ccurry cuncurry rule =
       
   167   let
       
   168     val cert = Thm.cterm_of (Proof_Context.theory_of ctxt)
       
   169     val ([P], ctxt') = Variable.variant_fixes ["P"] ctxt
       
   170 
       
   171     val split_paired_all_conv =
       
   172       Conv.every_conv (replicate (length args - 1) (Conv.rewr_conv @{thm split_paired_all}))
       
   173 
       
   174     val split_params_conv = 
       
   175       Conv.params_conv ~1 (fn ctxt' =>
       
   176         Conv.implies_conv split_paired_all_conv Conv.all_conv)
       
   177 
       
   178     val inst_rule =
       
   179       cterm_instantiate' [SOME cuncurry, NONE, SOME ccurry] rule
       
   180 
       
   181     val plain_resultT = 
       
   182       Thm.prop_of inst_rule |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop
       
   183       |> Term.head_of |> Term.dest_Var |> snd |> range_type |> domain_type
       
   184     val PT = map (snd o dest_Free) args ---> plain_resultT --> HOLogic.boolT
       
   185     val x_inst = cert (foldl1 HOLogic.mk_prod args)
       
   186     val P_inst = cert (uncurry_n (length args) (Free (P, PT)))
       
   187 
       
   188     val inst_rule' = inst_rule
       
   189       |> Tactic.rule_by_tactic ctxt
       
   190         (Simplifier.simp_tac curry_uncurry_ss 4
       
   191          THEN Simplifier.simp_tac curry_uncurry_ss 3
       
   192          THEN CONVERSION (split_params_conv ctxt
       
   193            then_conv (Conv.forall_conv (K split_paired_all_conv) ctxt)) 3)
       
   194       |> Drule.instantiate' [] [NONE, NONE, SOME P_inst, SOME x_inst]
       
   195       |> Simplifier.full_simplify split_conv_ss
       
   196       |> singleton (Variable.export ctxt' ctxt)
       
   197   in
       
   198     inst_rule'
       
   199   end;
       
   200     
   163 
   201 
   164 (*** partial_function definition ***)
   202 (*** partial_function definition ***)
   165 
   203 
   166 fun gen_add_partial_function prep mode fixes_raw eqn_raw lthy =
   204 fun gen_add_partial_function prep mode fixes_raw eqn_raw lthy =
   167   let
   205   let
   169       handle Option.Option => error (cat_lines ["Unknown mode " ^ quote mode ^ ".",
   207       handle Option.Option => error (cat_lines ["Unknown mode " ^ quote mode ^ ".",
   170         "Known modes are " ^ commas_quote (known_modes lthy) ^ "."]);
   208         "Known modes are " ^ commas_quote (known_modes lthy) ^ "."]);
   171     val Setup_Data {fixp, mono, fixp_eq, fixp_induct} = setup_data;
   209     val Setup_Data {fixp, mono, fixp_eq, fixp_induct} = setup_data;
   172 
   210 
   173     val ((fixes, [(eq_abinding, eqn)]), _) = prep fixes_raw [eqn_raw] lthy;
   211     val ((fixes, [(eq_abinding, eqn)]), _) = prep fixes_raw [eqn_raw] lthy;
   174     val ((_, plain_eqn), _) = Variable.focus eqn lthy;
   212     val ((_, plain_eqn), args_ctxt) = Variable.focus eqn lthy;
   175 
   213 
   176     val ((f_binding, fT), mixfix) = the_single fixes;
   214     val ((f_binding, fT), mixfix) = the_single fixes;
   177     val fname = Binding.name_of f_binding;
   215     val fname = Binding.name_of f_binding;
   178 
   216 
   179     val cert = cterm_of (Proof_Context.theory_of lthy);
   217     val cert = cterm_of (Proof_Context.theory_of lthy);
   180     val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop plain_eqn);
   218     val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop plain_eqn);
   181     val (head, args) = strip_comb lhs;
   219     val (head, args) = strip_comb lhs;
       
   220     val argnames = map (fst o dest_Free) args;
   182     val F = fold_rev lambda (head :: args) rhs;
   221     val F = fold_rev lambda (head :: args) rhs;
   183 
   222 
   184     val arity = length args;
   223     val arity = length args;
   185     val (aTs, bTs) = chop arity (binder_types fT);
   224     val (aTs, bTs) = chop arity (binder_types fT);
   186 
   225 
   214     val unfold =
   253     val unfold =
   215       (cterm_instantiate' (map (SOME o cert) [uncurry, F, curry]) fixp_eq
   254       (cterm_instantiate' (map (SOME o cert) [uncurry, F, curry]) fixp_eq
   216         OF [mono_thm, f_def])
   255         OF [mono_thm, f_def])
   217       |> Tactic.rule_by_tactic lthy (Simplifier.simp_tac curry_uncurry_ss 1);
   256       |> Tactic.rule_by_tactic lthy (Simplifier.simp_tac curry_uncurry_ss 1);
   218 
   257 
       
   258     val mk_raw_induct =
       
   259       mk_curried_induct args args_ctxt (cert curry) (cert uncurry)
       
   260       #> singleton (Variable.export args_ctxt lthy)
       
   261       #> (fn thm => Drule.instantiate' [] [SOME (cert F)] thm OF [mono_thm, f_def])
       
   262       #> Drule.rename_bvars' (map SOME (fname :: argnames @ argnames))
       
   263 
       
   264     val raw_induct = Option.map mk_raw_induct fixp_induct
   219     val rec_rule = let open Conv in
   265     val rec_rule = let open Conv in
   220       Goal.prove lthy' (map (fst o dest_Free) args) [] eqn (fn _ =>
   266       Goal.prove lthy' (map (fst o dest_Free) args) [] eqn (fn _ =>
   221         CONVERSION ((arg_conv o arg1_conv o head_conv o rewr_conv) (mk_meta_eq unfold)) 1
   267         CONVERSION ((arg_conv o arg1_conv o head_conv o rewr_conv) (mk_meta_eq unfold)) 1
   222         THEN rtac @{thm refl} 1) end;
   268         THEN rtac @{thm refl} 1) end;
   223   in
   269   in
   224     lthy'
   270     lthy'
   225     |> Local_Theory.note (eq_abinding, [rec_rule])
   271     |> Local_Theory.note (eq_abinding, [rec_rule])
   226     |-> (fn (_, rec') =>
   272     |-> (fn (_, rec') =>
   227       Spec_Rules.add Spec_Rules.Equational ([f], rec')
   273       Spec_Rules.add Spec_Rules.Equational ([f], rec')
   228       #> Local_Theory.note ((Binding.qualify true fname (Binding.name "simps"), []), rec') #> snd)
   274       #> Local_Theory.note ((Binding.qualify true fname (Binding.name "simps"), []), rec') #> snd)
       
   275     |> (case raw_induct of NONE => I | SOME thm =>
       
   276          Local_Theory.note ((Binding.qualify true fname (Binding.name "raw_induct"), []), [thm]) #> snd)
   229   end;
   277   end;
   230 
   278 
   231 val add_partial_function = gen_add_partial_function Specification.check_spec;
   279 val add_partial_function = gen_add_partial_function Specification.check_spec;
   232 val add_partial_function_cmd = gen_add_partial_function Specification.read_spec;
   280 val add_partial_function_cmd = gen_add_partial_function Specification.read_spec;
   233 
   281