| author | wenzelm | 
| Sun, 23 Mar 2025 17:07:55 +0100 | |
| changeset 82322 | 94fd80f0107d | 
| parent 81946 | ee680c69de38 | 
| child 82641 | d22294b20573 | 
| permissions | -rw-r--r-- | 
| 40107 | 1 | (* Title: HOL/Tools/Function/partial_function.ML | 
| 2 | Author: Alexander Krauss, TU Muenchen | |
| 3 | ||
| 4 | Partial function definitions based on least fixed points in ccpos. | |
| 5 | *) | |
| 6 | ||
| 7 | signature PARTIAL_FUNCTION = | |
| 8 | sig | |
| 78085 | 9 | val init: string -> term -> term -> thm -> thm -> thm option -> Morphism.declaration | 
| 52727 | 10 | val mono_tac: Proof.context -> int -> tactic | 
| 40107 | 11 | val add_partial_function: string -> (binding * typ option * mixfix) list -> | 
| 66653 | 12 | Attrib.binding * term -> local_theory -> (term * thm) * local_theory | 
| 40107 | 13 | val add_partial_function_cmd: string -> (binding * string option * mixfix) list -> | 
| 66653 | 14 | Attrib.binding * string -> local_theory -> (term * thm) * local_theory | 
| 69829 | 15 | val transform_result: morphism -> term * thm -> term * thm | 
| 40107 | 16 | end; | 
| 17 | ||
| 18 | structure Partial_Function: PARTIAL_FUNCTION = | |
| 19 | struct | |
| 20 | ||
| 63005 | 21 | open Function_Lib | 
| 22 | ||
| 23 | ||
| 40107 | 24 | (*** Context Data ***) | 
| 25 | ||
| 61113 | 26 | datatype setup_data = Setup_Data of | 
| 43080 | 27 |  {fixp: term,
 | 
| 28 | mono: term, | |
| 29 | fixp_eq: thm, | |
| 52728 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 krauss parents: 
52727diff
changeset | 30 | fixp_induct: thm, | 
| 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 krauss parents: 
52727diff
changeset | 31 | fixp_induct_user: thm option}; | 
| 43080 | 32 | |
| 61113 | 33 | fun transform_setup_data phi (Setup_Data {fixp, mono, fixp_eq, fixp_induct, fixp_induct_user}) =
 | 
| 34 | let | |
| 35 | val term = Morphism.term phi; | |
| 36 | val thm = Morphism.thm phi; | |
| 37 | in | |
| 38 | Setup_Data | |
| 39 |      {fixp = term fixp, mono = term mono, fixp_eq = thm fixp_eq,
 | |
| 40 | fixp_induct = thm fixp_induct, fixp_induct_user = Option.map thm fixp_induct_user} | |
| 41 | end; | |
| 42 | ||
| 40107 | 43 | structure Modes = Generic_Data | 
| 44 | ( | |
| 43080 | 45 | type T = setup_data Symtab.table; | 
| 40107 | 46 | val empty = Symtab.empty; | 
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
41117diff
changeset | 47 | fun merge data = Symtab.merge (K true) data; | 
| 40107 | 48 | ) | 
| 49 | ||
| 52728 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 krauss parents: 
52727diff
changeset | 50 | fun init mode fixp mono fixp_eq fixp_induct fixp_induct_user phi = | 
| 40107 | 51 | let | 
| 61113 | 52 | val data' = | 
| 53 | Setup_Data | |
| 54 |        {fixp = fixp, mono = mono, fixp_eq = fixp_eq,
 | |
| 55 | fixp_induct = fixp_induct, fixp_induct_user = fixp_induct_user} | |
| 56 | |> transform_setup_data (phi $> Morphism.trim_context_morphism); | |
| 57 | in Modes.map (Symtab.update (mode, data')) end; | |
| 40107 | 58 | |
| 59 | val known_modes = Symtab.keys o Modes.get o Context.Proof; | |
| 61113 | 60 | |
| 61 | fun lookup_mode ctxt = | |
| 62 | Symtab.lookup (Modes.get (Context.Proof ctxt)) | |
| 67664 | 63 | #> Option.map (transform_setup_data (Morphism.transfer_morphism' ctxt)); | 
| 40107 | 64 | |
| 65 | ||
| 66 | (*** Automated monotonicity proofs ***) | |
| 67 | ||
| 68 | (*rewrite conclusion with k-th assumtion*) | |
| 69 | fun rewrite_with_asm_tac ctxt k = | |
| 45403 | 70 |   Subgoal.FOCUS (fn {context = ctxt', prems, ...} =>
 | 
| 63170 | 71 | Local_Defs.unfold0_tac ctxt' [nth prems k]) ctxt; | 
| 40107 | 72 | |
| 54405 
88f6d5b1422f
ported 'partial_function' to 'Ctr_Sugar' abstraction
 blanchet parents: 
52728diff
changeset | 73 | fun dest_case ctxt t = | 
| 40107 | 74 | case strip_comb t of | 
| 75 | (Const (case_comb, _), args) => | |
| 54405 
88f6d5b1422f
ported 'partial_function' to 'Ctr_Sugar' abstraction
 blanchet parents: 
52728diff
changeset | 76 | (case Ctr_Sugar.ctr_sugar_of_case ctxt case_comb of | 
| 40107 | 77 | NONE => NONE | 
| 54405 
88f6d5b1422f
ported 'partial_function' to 'Ctr_Sugar' abstraction
 blanchet parents: 
52728diff
changeset | 78 |        | SOME {case_thms, ...} =>
 | 
| 40107 | 79 | let | 
| 59582 | 80 | val lhs = Thm.prop_of (hd case_thms) | 
| 40107 | 81 | |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst; | 
| 82 | val arity = length (snd (strip_comb lhs)); | |
| 83 | val conv = funpow (length args - arity) Conv.fun_conv | |
| 54405 
88f6d5b1422f
ported 'partial_function' to 'Ctr_Sugar' abstraction
 blanchet parents: 
52728diff
changeset | 84 | (Conv.rewrs_conv (map mk_meta_eq case_thms)); | 
| 40107 | 85 | in | 
| 86 | SOME (nth args (arity - 1), conv) | |
| 87 | end) | |
| 88 | | _ => NONE; | |
| 89 | ||
| 90 | (*split on case expressions*) | |
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58839diff
changeset | 91 | val split_cases_tac = Subgoal.FOCUS_PARAMS (fn {context = ctxt, ...} =>
 | 
| 40107 | 92 | SUBGOAL (fn (t, i) => case t of | 
| 93 | _ $ (_ $ Abs (_, _, body)) => | |
| 54405 
88f6d5b1422f
ported 'partial_function' to 'Ctr_Sugar' abstraction
 blanchet parents: 
52728diff
changeset | 94 | (case dest_case ctxt body of | 
| 40107 | 95 | NONE => no_tac | 
| 96 | | SOME (arg, conv) => | |
| 97 | let open Conv in | |
| 42083 
e1209fc7ecdc
added Term.is_open and Term.is_dependent convenience, to cover common situations of loose bounds;
 wenzelm parents: 
41472diff
changeset | 98 | if Term.is_open arg then no_tac | 
| 61844 | 99 | else ((DETERM o Induct.cases_tac ctxt false [[SOME arg]] NONE []) | 
| 40107 | 100 | THEN_ALL_NEW (rewrite_with_asm_tac ctxt 0) | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58839diff
changeset | 101 |                 THEN_ALL_NEW eresolve_tac ctxt @{thms thin_rl}
 | 
| 40107 | 102 | THEN_ALL_NEW (CONVERSION | 
| 103 | (params_conv ~1 (fn ctxt' => | |
| 104 | arg_conv (arg_conv (abs_conv (K conv) ctxt'))) ctxt))) i | |
| 105 | end) | |
| 106 | | _ => no_tac) 1); | |
| 107 | ||
| 108 | (*monotonicity proof: apply rules + split case expressions*) | |
| 109 | fun mono_tac ctxt = | |
| 63170 | 110 |   K (Local_Defs.unfold0_tac ctxt [@{thm curry_def}])
 | 
| 40107 | 111 | THEN' (TRY o REPEAT_ALL_NEW | 
| 67149 | 112 | (resolve_tac ctxt (rev (Named_Theorems.get ctxt \<^named_theorems>\<open>partial_function_mono\<close>)) | 
| 40107 | 113 | ORELSE' split_cases_tac ctxt)); | 
| 114 | ||
| 115 | ||
| 116 | (*** Auxiliary functions ***) | |
| 117 | ||
| 118 | (*Returns t $ u, but instantiates the type of t to make the | |
| 119 | application type correct*) | |
| 120 | fun apply_inst ctxt t u = | |
| 121 | let | |
| 42361 | 122 | val thy = Proof_Context.theory_of ctxt; | 
| 40107 | 123 | val T = domain_type (fastype_of t); | 
| 124 | val T' = fastype_of u; | |
| 42388 | 125 | val subst = Sign.typ_match thy (T, T') Vartab.empty | 
| 40107 | 126 |       handle Type.TYPE_MATCH => raise TYPE ("apply_inst", [T, T'], [t, u])
 | 
| 127 | in | |
| 128 | map_types (Envir.norm_type subst) t $ u | |
| 129 | end; | |
| 130 | ||
| 131 | fun head_conv cv ct = | |
| 81946 
ee680c69de38
misc tuning: prefer specific variants of Thm.dest_comb;
 wenzelm parents: 
78085diff
changeset | 132 | if can Thm.dest_fun ct then Conv.fun_conv (head_conv cv) ct else cv ct; | 
| 40107 | 133 | |
| 134 | ||
| 135 | (*** currying transformation ***) | |
| 136 | ||
| 137 | fun curry_const (A, B, C) = | |
| 67149 | 138 | Const (\<^const_name>\<open>Product_Type.curry\<close>, | 
| 40107 | 139 | [HOLogic.mk_prodT (A, B) --> C, A, B] ---> C); | 
| 140 | ||
| 141 | fun mk_curry f = | |
| 142 | case fastype_of f of | |
| 143 |     Type ("fun", [Type (_, [S, T]), U]) =>
 | |
| 144 | curry_const (S, T, U) $ f | |
| 145 |   | T => raise TYPE ("mk_curry", [T], [f]);
 | |
| 146 | ||
| 147 | (* iterated versions. Nonstandard left-nested tuples arise naturally | |
| 148 | from "split o split o split"*) | |
| 149 | fun curry_n arity = funpow (arity - 1) mk_curry; | |
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61121diff
changeset | 150 | fun uncurry_n arity = funpow (arity - 1) HOLogic.mk_case_prod; | 
| 40107 | 151 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51484diff
changeset | 152 | val curry_uncurry_ss = | 
| 69593 | 153 | simpset_of (put_simpset HOL_basic_ss \<^context> | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61121diff
changeset | 154 |     addsimps [@{thm Product_Type.curry_case_prod}, @{thm Product_Type.case_prod_curry}])
 | 
| 40107 | 155 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51484diff
changeset | 156 | val split_conv_ss = | 
| 69593 | 157 | simpset_of (put_simpset HOL_basic_ss \<^context> | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51484diff
changeset | 158 |     addsimps [@{thm Product_Type.split_conv}]);
 | 
| 43083 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 krauss parents: 
43080diff
changeset | 159 | |
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
54405diff
changeset | 160 | val curry_K_ss = | 
| 69593 | 161 | simpset_of (put_simpset HOL_basic_ss \<^context> | 
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
54405diff
changeset | 162 |     addsimps [@{thm Product_Type.curry_K}]);
 | 
| 52728 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 krauss parents: 
52727diff
changeset | 163 | |
| 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 krauss parents: 
52727diff
changeset | 164 | (* instantiate generic fixpoint induction and eliminate the canonical assumptions; | 
| 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 krauss parents: 
52727diff
changeset | 165 | curry induction predicate *) | 
| 70325 | 166 | fun specialize_fixp_induct ctxt fT fT_uc curry uncurry mono_thm f_def rule = | 
| 52728 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 krauss parents: 
52727diff
changeset | 167 | let | 
| 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 krauss parents: 
52727diff
changeset | 168 | val ([P], ctxt') = Variable.variant_fixes ["P"] ctxt | 
| 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 krauss parents: 
52727diff
changeset | 169 |     val P_inst = Abs ("f", fT_uc, Free (P, fT --> HOLogic.boolT) $ (curry $ Bound 0))
 | 
| 61113 | 170 | in | 
| 70324 
005c1cdbfd3f
backout 34bc296374ee -- affects the raw_induct rule, e.g. relevant for AFP/Imperative_Insertion_Sort;
 wenzelm parents: 
70319diff
changeset | 171 | (* FIXME ctxt vs. ctxt' (!?) *) | 
| 52728 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 krauss parents: 
52727diff
changeset | 172 | rule | 
| 70324 
005c1cdbfd3f
backout 34bc296374ee -- affects the raw_induct rule, e.g. relevant for AFP/Imperative_Insertion_Sort;
 wenzelm parents: 
70319diff
changeset | 173 | |> infer_instantiate' ctxt | 
| 
005c1cdbfd3f
backout 34bc296374ee -- affects the raw_induct rule, e.g. relevant for AFP/Imperative_Insertion_Sort;
 wenzelm parents: 
70319diff
changeset | 174 | ((map o Option.map) (Thm.cterm_of ctxt) [SOME uncurry, NONE, SOME curry, NONE, SOME P_inst]) | 
| 
005c1cdbfd3f
backout 34bc296374ee -- affects the raw_induct rule, e.g. relevant for AFP/Imperative_Insertion_Sort;
 wenzelm parents: 
70319diff
changeset | 175 | |> Tactic.rule_by_tactic ctxt | 
| 
005c1cdbfd3f
backout 34bc296374ee -- affects the raw_induct rule, e.g. relevant for AFP/Imperative_Insertion_Sort;
 wenzelm parents: 
70319diff
changeset | 176 | (Simplifier.simp_tac (put_simpset curry_uncurry_ss ctxt) 3 (* discharge U (C f) = f *) | 
| 
005c1cdbfd3f
backout 34bc296374ee -- affects the raw_induct rule, e.g. relevant for AFP/Imperative_Insertion_Sort;
 wenzelm parents: 
70319diff
changeset | 177 | THEN Simplifier.simp_tac (put_simpset curry_K_ss ctxt) 4 (* simplify bot case *) | 
| 
005c1cdbfd3f
backout 34bc296374ee -- affects the raw_induct rule, e.g. relevant for AFP/Imperative_Insertion_Sort;
 wenzelm parents: 
70319diff
changeset | 178 | THEN Simplifier.full_simp_tac (put_simpset curry_uncurry_ss ctxt) 5) (* simplify induction step *) | 
| 52728 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 krauss parents: 
52727diff
changeset | 179 | |> (fn thm => thm OF [mono_thm, f_def]) | 
| 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 krauss parents: 
52727diff
changeset | 180 | |> Conv.fconv_rule (Conv.concl_conv ~1 (* simplify conclusion *) | 
| 70324 
005c1cdbfd3f
backout 34bc296374ee -- affects the raw_induct rule, e.g. relevant for AFP/Imperative_Insertion_Sort;
 wenzelm parents: 
70319diff
changeset | 181 |          (Raw_Simplifier.rewrite ctxt false [mk_meta_eq @{thm Product_Type.curry_case_prod}]))
 | 
| 52728 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 krauss parents: 
52727diff
changeset | 182 | |> singleton (Variable.export ctxt' ctxt) | 
| 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 krauss parents: 
52727diff
changeset | 183 | end | 
| 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 krauss parents: 
52727diff
changeset | 184 | |
| 52727 | 185 | fun mk_curried_induct args ctxt inst_rule = | 
| 43083 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 krauss parents: 
43080diff
changeset | 186 | let | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59582diff
changeset | 187 | val cert = Thm.cterm_of ctxt | 
| 70325 | 188 | (* FIXME ctxt vs. ctxt' (!?) *) | 
| 43083 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 krauss parents: 
43080diff
changeset | 189 | val ([P], ctxt') = Variable.variant_fixes ["P"] ctxt | 
| 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 krauss parents: 
43080diff
changeset | 190 | |
| 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 krauss parents: 
43080diff
changeset | 191 | val split_paired_all_conv = | 
| 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 krauss parents: 
43080diff
changeset | 192 |       Conv.every_conv (replicate (length args - 1) (Conv.rewr_conv @{thm split_paired_all}))
 | 
| 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 krauss parents: 
43080diff
changeset | 193 | |
| 61113 | 194 | val split_params_conv = | 
| 70325 | 195 | Conv.params_conv ~1 (fn _ => Conv.implies_conv split_paired_all_conv Conv.all_conv) | 
| 43083 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 krauss parents: 
43080diff
changeset | 196 | |
| 52521 
a1c4f586e372
more robust instantiation: extract vars from conclusion instead of relying on fixed overall ordering
 krauss parents: 
52087diff
changeset | 197 | val (P_var, x_var) = | 
| 51484 
49eb8d73ae10
allow induction predicates with arbitrary arity (not just binary)
 krauss parents: 
46961diff
changeset | 198 | Thm.prop_of inst_rule |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop | 
| 52521 
a1c4f586e372
more robust instantiation: extract vars from conclusion instead of relying on fixed overall ordering
 krauss parents: 
52087diff
changeset | 199 | |> strip_comb |> apsnd hd | 
| 60642 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
59936diff
changeset | 200 | |> apply2 dest_Var | 
| 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
59936diff
changeset | 201 | val P_rangeT = range_type (snd P_var) | 
| 51484 
49eb8d73ae10
allow induction predicates with arbitrary arity (not just binary)
 krauss parents: 
46961diff
changeset | 202 | val PT = map (snd o dest_Free) args ---> P_rangeT | 
| 43083 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 krauss parents: 
43080diff
changeset | 203 | val x_inst = cert (foldl1 HOLogic.mk_prod args) | 
| 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 krauss parents: 
43080diff
changeset | 204 | val P_inst = cert (uncurry_n (length args) (Free (P, PT))) | 
| 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 krauss parents: 
43080diff
changeset | 205 | |
| 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 krauss parents: 
43080diff
changeset | 206 | val inst_rule' = inst_rule | 
| 70324 
005c1cdbfd3f
backout 34bc296374ee -- affects the raw_induct rule, e.g. relevant for AFP/Imperative_Insertion_Sort;
 wenzelm parents: 
70319diff
changeset | 207 | |> Tactic.rule_by_tactic ctxt | 
| 
005c1cdbfd3f
backout 34bc296374ee -- affects the raw_induct rule, e.g. relevant for AFP/Imperative_Insertion_Sort;
 wenzelm parents: 
70319diff
changeset | 208 | (Simplifier.simp_tac (put_simpset curry_uncurry_ss ctxt) 4 | 
| 
005c1cdbfd3f
backout 34bc296374ee -- affects the raw_induct rule, e.g. relevant for AFP/Imperative_Insertion_Sort;
 wenzelm parents: 
70319diff
changeset | 209 | THEN Simplifier.simp_tac (put_simpset curry_uncurry_ss ctxt) 3 | 
| 
005c1cdbfd3f
backout 34bc296374ee -- affects the raw_induct rule, e.g. relevant for AFP/Imperative_Insertion_Sort;
 wenzelm parents: 
70319diff
changeset | 210 | THEN CONVERSION (split_params_conv ctxt | 
| 
005c1cdbfd3f
backout 34bc296374ee -- affects the raw_induct rule, e.g. relevant for AFP/Imperative_Insertion_Sort;
 wenzelm parents: 
70319diff
changeset | 211 | then_conv (Conv.forall_conv (K split_paired_all_conv) ctxt)) 3) | 
| 77879 | 212 | |> Thm.instantiate (TVars.empty, Vars.make2 (P_var, P_inst) (x_var, x_inst)) | 
| 70324 
005c1cdbfd3f
backout 34bc296374ee -- affects the raw_induct rule, e.g. relevant for AFP/Imperative_Insertion_Sort;
 wenzelm parents: 
70319diff
changeset | 213 | |> Simplifier.full_simplify (put_simpset split_conv_ss ctxt) | 
| 43083 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 krauss parents: 
43080diff
changeset | 214 | |> singleton (Variable.export ctxt' ctxt) | 
| 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 krauss parents: 
43080diff
changeset | 215 | in | 
| 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 krauss parents: 
43080diff
changeset | 216 | inst_rule' | 
| 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 krauss parents: 
43080diff
changeset | 217 | end; | 
| 61113 | 218 | |
| 40107 | 219 | |
| 220 | (*** partial_function definition ***) | |
| 221 | ||
| 69829 | 222 | fun transform_result phi (t, thm) = (Morphism.term phi t, Morphism.thm phi thm); | 
| 223 | ||
| 40107 | 224 | fun gen_add_partial_function prep mode fixes_raw eqn_raw lthy = | 
| 225 | let | |
| 43080 | 226 | val setup_data = the (lookup_mode lthy mode) | 
| 40107 | 227 | handle Option.Option => error (cat_lines ["Unknown mode " ^ quote mode ^ ".", | 
| 228 | "Known modes are " ^ commas_quote (known_modes lthy) ^ "."]); | |
| 52728 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 krauss parents: 
52727diff
changeset | 229 |     val Setup_Data {fixp, mono, fixp_eq, fixp_induct, fixp_induct_user} = setup_data;
 | 
| 40107 | 230 | |
| 63182 | 231 | val ((fixes, [(eq_abinding, eqn)]), _) = prep fixes_raw [(eqn_raw, [], [])] lthy; | 
| 60695 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 wenzelm parents: 
60642diff
changeset | 232 | val ((_, plain_eqn), args_ctxt) = Variable.focus NONE eqn lthy; | 
| 40107 | 233 | |
| 234 | val ((f_binding, fT), mixfix) = the_single fixes; | |
| 62997 | 235 | val f_bname = Binding.name_of f_binding; | 
| 40107 | 236 | |
| 62998 | 237 | fun note_qualified (name, thms) = | 
| 63005 | 238 | Local_Theory.note ((derived_name f_binding name, []), thms) #> snd | 
| 62998 | 239 | |
| 40107 | 240 | val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop plain_eqn); | 
| 241 | val (head, args) = strip_comb lhs; | |
| 43083 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 krauss parents: 
43080diff
changeset | 242 | val argnames = map (fst o dest_Free) args; | 
| 40107 | 243 | val F = fold_rev lambda (head :: args) rhs; | 
| 244 | ||
| 245 | val arity = length args; | |
| 246 | val (aTs, bTs) = chop arity (binder_types fT); | |
| 247 | ||
| 248 | val tupleT = foldl1 HOLogic.mk_prodT aTs; | |
| 249 | val fT_uc = tupleT :: bTs ---> body_type fT; | |
| 62997 | 250 | val f_uc = Var ((f_bname, 0), fT_uc); | 
| 63008 | 251 |     val x_uc = Var (("x", 1), tupleT);
 | 
| 40107 | 252 | val uncurry = lambda head (uncurry_n arity head); | 
| 253 | val curry = lambda f_uc (curry_n arity f_uc); | |
| 254 | ||
| 255 | val F_uc = | |
| 256 | lambda f_uc (uncurry_n arity (F $ curry_n arity f_uc)); | |
| 257 | ||
| 258 | val mono_goal = apply_inst lthy mono (lambda f_uc (F_uc $ f_uc $ x_uc)) | |
| 259 | |> HOLogic.mk_Trueprop | |
| 260 | |> Logic.all x_uc; | |
| 261 | ||
| 74530 
823ccd84b879
revert bbfed17243af, breaks HOL-Proofs extraction;
 wenzelm parents: 
74526diff
changeset | 262 | val mono_thm = Goal.prove_internal lthy [] (Thm.cterm_of lthy mono_goal) | 
| 
823ccd84b879
revert bbfed17243af, breaks HOL-Proofs extraction;
 wenzelm parents: 
74526diff
changeset | 263 | (K (mono_tac lthy 1)) | 
| 60784 | 264 | val inst_mono_thm = Thm.forall_elim (Thm.cterm_of lthy x_uc) mono_thm | 
| 40107 | 265 | |
| 266 | val f_def_rhs = curry_n arity (apply_inst lthy fixp F_uc); | |
| 61121 
efe8b18306b7
do not expose low-level "_def" facts of 'function' definitions, to avoid potential confusion with the situation of plain 'definition';
 wenzelm parents: 
61113diff
changeset | 267 | val f_def_binding = | 
| 62997 | 268 | Thm.make_def_binding (Config.get lthy Function_Lib.function_internals) f_binding | 
| 269 | val ((f, (_, f_def)), lthy') = | |
| 270 | Local_Theory.define ((f_binding, mixfix), ((f_def_binding, []), f_def_rhs)) lthy; | |
| 40107 | 271 | |
| 272 | val eqn = HOLogic.mk_eq (list_comb (f, args), | |
| 273 | Term.betapplys (F, f :: args)) | |
| 274 | |> HOLogic.mk_Trueprop; | |
| 275 | ||
| 276 | val unfold = | |
| 60784 | 277 | (infer_instantiate' lthy' (map (SOME o Thm.cterm_of lthy') [uncurry, F, curry]) fixp_eq | 
| 52727 | 278 | OF [inst_mono_thm, f_def]) | 
| 52087 
f3075fc4f5f6
more precise treatment of theory vs. Proof.context;
 wenzelm parents: 
51717diff
changeset | 279 | |> Tactic.rule_by_tactic lthy' (Simplifier.simp_tac (put_simpset curry_uncurry_ss lthy') 1); | 
| 40107 | 280 | |
| 61113 | 281 | val specialized_fixp_induct = | 
| 70325 | 282 | specialize_fixp_induct lthy' fT fT_uc curry uncurry inst_mono_thm f_def fixp_induct | 
| 62997 | 283 | |> Drule.rename_bvars' (map SOME (f_bname :: f_bname :: argnames)); | 
| 52728 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 krauss parents: 
52727diff
changeset | 284 | |
| 43083 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 krauss parents: 
43080diff
changeset | 285 | val mk_raw_induct = | 
| 60784 | 286 | infer_instantiate' args_ctxt | 
| 287 | ((map o Option.map) (Thm.cterm_of args_ctxt) [SOME uncurry, NONE, SOME curry]) | |
| 52727 | 288 | #> mk_curried_induct args args_ctxt | 
| 52087 
f3075fc4f5f6
more precise treatment of theory vs. Proof.context;
 wenzelm parents: 
51717diff
changeset | 289 | #> singleton (Variable.export args_ctxt lthy') | 
| 60784 | 290 | #> (fn thm => infer_instantiate' lthy' | 
| 291 | [SOME (Thm.cterm_of lthy' F)] thm OF [inst_mono_thm, f_def]) | |
| 62997 | 292 | #> Drule.rename_bvars' (map SOME (f_bname :: argnames @ argnames)) | 
| 43083 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 krauss parents: 
43080diff
changeset | 293 | |
| 52728 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 krauss parents: 
52727diff
changeset | 294 | val raw_induct = Option.map mk_raw_induct fixp_induct_user | 
| 62997 | 295 | val rec_rule = | 
| 296 | let open Conv in | |
| 297 | Goal.prove lthy' (map (fst o dest_Free) args) [] eqn (fn _ => | |
| 298 | CONVERSION ((arg_conv o arg1_conv o head_conv o rewr_conv) (mk_meta_eq unfold)) 1 | |
| 299 |           THEN resolve_tac lthy' @{thms refl} 1)
 | |
| 300 | end; | |
| 66653 | 301 | val ((_, [rec_rule']), lthy'') = lthy' |> Local_Theory.note (eq_abinding, [rec_rule]) | 
| 40107 | 302 | in | 
| 66653 | 303 | lthy'' | 
| 71214 
5727bcc3c47c
proper spec_rule name via naming/binding/Morphism.binding;
 wenzelm parents: 
71179diff
changeset | 304 | |> Spec_Rules.add Binding.empty Spec_Rules.equational_recdef [f] [rec_rule'] | 
| 66653 | 305 |     |> note_qualified ("simps", [rec_rule'])
 | 
| 62998 | 306 |     |> note_qualified ("mono", [mono_thm])
 | 
| 307 |     |> (case raw_induct of NONE => I | SOME thm => note_qualified ("raw_induct", [thm]))
 | |
| 308 |     |> note_qualified ("fixp_induct", [specialized_fixp_induct])
 | |
| 66653 | 309 | |> pair (f, rec_rule') | 
| 40107 | 310 | end; | 
| 311 | ||
| 63064 
2f18172214c8
support 'assumes' in specifications, e.g. 'definition', 'inductive';
 wenzelm parents: 
63008diff
changeset | 312 | val add_partial_function = gen_add_partial_function Specification.check_multi_specs; | 
| 
2f18172214c8
support 'assumes' in specifications, e.g. 'definition', 'inductive';
 wenzelm parents: 
63008diff
changeset | 313 | val add_partial_function_cmd = gen_add_partial_function Specification.read_multi_specs; | 
| 40107 | 314 | |
| 67149 | 315 | val mode = \<^keyword>\<open>(\<close> |-- Parse.name --| \<^keyword>\<open>)\<close>; | 
| 40107 | 316 | |
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46949diff
changeset | 317 | val _ = | 
| 67149 | 318 | Outer_Syntax.local_theory \<^command_keyword>\<open>partial_function\<close> "define partial function" | 
| 63285 | 319 | ((mode -- (Parse.vars -- (Parse.where_ |-- Parse_Spec.simple_spec))) | 
| 66653 | 320 | >> (fn (mode, (vars, spec)) => add_partial_function_cmd mode vars spec #> #2)); | 
| 40107 | 321 | |
| 61113 | 322 | end; |