src/HOL/Tools/Function/function_lib.ML
author haftmann
Fri, 20 Oct 2017 20:57:55 +0200
changeset 66888 930abfdf8727
parent 63011 301e631666a0
child 67149 e61557884799
permissions -rw-r--r--
algebraic foundation for congruences
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37744
3daaf23b9ab4 tuned titles
haftmann
parents: 37391
diff changeset
     1
(*  Title:      HOL/Tools/Function/function_lib.ML
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     2
    Author:     Alexander Krauss, TU Muenchen
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     3
41113
b223fa19af3c added signature;
krauss
parents: 40722
diff changeset
     4
Ad-hoc collection of function waiting to be eliminated, generalized,
b223fa19af3c added signature;
krauss
parents: 40722
diff changeset
     5
moved elsewhere or otherwise cleaned up.
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     6
*)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     7
41113
b223fa19af3c added signature;
krauss
parents: 40722
diff changeset
     8
signature FUNCTION_LIB =
b223fa19af3c added signature;
krauss
parents: 40722
diff changeset
     9
sig
62093
bd73a2279fcd more uniform treatment of package internals;
wenzelm
parents: 61121
diff changeset
    10
  val function_internals: bool Config.T
61121
efe8b18306b7 do not expose low-level "_def" facts of 'function' definitions, to avoid potential confusion with the situation of plain 'definition';
wenzelm
parents: 60328
diff changeset
    11
63005
d743bb1b6c23 clarified bindings;
wenzelm
parents: 62093
diff changeset
    12
  val derived_name: binding -> string -> binding
d743bb1b6c23 clarified bindings;
wenzelm
parents: 62093
diff changeset
    13
  val derived_name_suffix: binding -> string -> binding
d743bb1b6c23 clarified bindings;
wenzelm
parents: 62093
diff changeset
    14
41113
b223fa19af3c added signature;
krauss
parents: 40722
diff changeset
    15
  val plural: string -> string -> 'a list -> string
b223fa19af3c added signature;
krauss
parents: 40722
diff changeset
    16
b223fa19af3c added signature;
krauss
parents: 40722
diff changeset
    17
  val dest_all_all: term -> term list * term
b223fa19af3c added signature;
krauss
parents: 40722
diff changeset
    18
b223fa19af3c added signature;
krauss
parents: 40722
diff changeset
    19
  val unordered_pairs: 'a list -> ('a * 'a) list
b223fa19af3c added signature;
krauss
parents: 40722
diff changeset
    20
b223fa19af3c added signature;
krauss
parents: 40722
diff changeset
    21
  val rename_bound: string -> term -> term
b223fa19af3c added signature;
krauss
parents: 40722
diff changeset
    22
  val mk_forall_rename: (string * term) -> term -> term
b223fa19af3c added signature;
krauss
parents: 40722
diff changeset
    23
  val forall_intr_rename: (string * cterm) -> thm -> thm
b223fa19af3c added signature;
krauss
parents: 40722
diff changeset
    24
b223fa19af3c added signature;
krauss
parents: 40722
diff changeset
    25
  datatype proof_attempt = Solved of thm | Stuck of thm | Fail
60328
9c94e6a30d29 clarified context;
wenzelm
parents: 59625
diff changeset
    26
  val try_proof: Proof.context -> cterm -> tactic -> proof_attempt
41113
b223fa19af3c added signature;
krauss
parents: 40722
diff changeset
    27
b223fa19af3c added signature;
krauss
parents: 40722
diff changeset
    28
  val dest_binop_list: string -> term -> term list
59625
aacdce52b2fc proper context;
wenzelm
parents: 59621
diff changeset
    29
  val regroup_conv: Proof.context -> string -> string -> thm list -> int list -> conv
aacdce52b2fc proper context;
wenzelm
parents: 59621
diff changeset
    30
  val regroup_union_conv: Proof.context -> int list -> conv
54406
a2d18fea844a undid copy-paste
blanchet
parents: 45156
diff changeset
    31
54407
e95831757903 ported part of function package to new 'Ctr_Sugar' abstraction
blanchet
parents: 54406
diff changeset
    32
  val inst_constrs_of: Proof.context -> typ -> term list
41113
b223fa19af3c added signature;
krauss
parents: 40722
diff changeset
    33
end
b223fa19af3c added signature;
krauss
parents: 40722
diff changeset
    34
b223fa19af3c added signature;
krauss
parents: 40722
diff changeset
    35
structure Function_Lib: FUNCTION_LIB =
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    36
struct
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    37
62093
bd73a2279fcd more uniform treatment of package internals;
wenzelm
parents: 61121
diff changeset
    38
val function_internals = Attrib.setup_config_bool @{binding function_internals} (K false)
61121
efe8b18306b7 do not expose low-level "_def" facts of 'function' definitions, to avoid potential confusion with the situation of plain 'definition';
wenzelm
parents: 60328
diff changeset
    39
63005
d743bb1b6c23 clarified bindings;
wenzelm
parents: 62093
diff changeset
    40
fun derived_name binding name =
d743bb1b6c23 clarified bindings;
wenzelm
parents: 62093
diff changeset
    41
  Binding.reset_pos (Binding.qualify_name true binding name)
d743bb1b6c23 clarified bindings;
wenzelm
parents: 62093
diff changeset
    42
63011
301e631666a0 clarified bindings;
wenzelm
parents: 63005
diff changeset
    43
fun derived_name_suffix binding sffx =
301e631666a0 clarified bindings;
wenzelm
parents: 63005
diff changeset
    44
  Binding.reset_pos (Binding.map_name (suffix sffx) binding)
63005
d743bb1b6c23 clarified bindings;
wenzelm
parents: 62093
diff changeset
    45
61121
efe8b18306b7 do not expose low-level "_def" facts of 'function' definitions, to avoid potential confusion with the situation of plain 'definition';
wenzelm
parents: 60328
diff changeset
    46
40076
6f012a209dac some cleanup in Function_Lib
krauss
parents: 39756
diff changeset
    47
(* "The variable" ^ plural " is" "s are" vs *)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    48
fun plural sg pl [x] = sg
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    49
  | plural sg pl _ = pl
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    50
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    51
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    52
(* Removes all quantifiers from a term, replacing bound variables by frees. *)
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 54883
diff changeset
    53
fun dest_all_all (t as (Const (@{const_name Pure.all},_) $ _)) =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    54
  let
40076
6f012a209dac some cleanup in Function_Lib
krauss
parents: 39756
diff changeset
    55
    val (v,b) = Logic.dest_all t |> apfst Free
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    56
    val (vs, b') = dest_all_all b
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    57
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    58
    (v :: vs, b')
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    59
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    60
  | dest_all_all t = ([],t)
33611
haftmann
parents: 33099
diff changeset
    61
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    62
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    63
(* forms all "unordered pairs": [1, 2, 3] ==> [(1, 1), (1, 2), (1, 3), (2, 2), (2, 3), (3, 3)] *)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    64
fun unordered_pairs [] = []
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    65
  | unordered_pairs (x::xs) = map (pair x) (x::xs) @ unordered_pairs xs
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    66
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    67
42497
89acb654d8eb inlined Function_Lib.replace_frees, which is used only once
krauss
parents: 42483
diff changeset
    68
(* renaming bound variables *)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    69
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    70
fun rename_bound n (Q $ Abs (_, T, b)) = (Q $ Abs (n, T, b))
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    71
  | rename_bound n _ = raise Match
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    72
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    73
fun mk_forall_rename (n, v) =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    74
  rename_bound n o Logic.all v
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    75
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    76
fun forall_intr_rename (n, cv) thm =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    77
  let
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 35402
diff changeset
    78
    val allthm = Thm.forall_intr cv thm
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    79
    val (_ $ abs) = Thm.prop_of allthm
45156
a9b6c2ea7bec added Term.dummy_pattern conveniences;
wenzelm
parents: 42498
diff changeset
    80
  in Thm.rename_boundvars abs (Abs (n, dummyT, Term.dummy)) allthm end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    81
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    82
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    83
datatype proof_attempt = Solved of thm | Stuck of thm | Fail
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    84
60328
9c94e6a30d29 clarified context;
wenzelm
parents: 59625
diff changeset
    85
fun try_proof ctxt cgoal tac =
9c94e6a30d29 clarified context;
wenzelm
parents: 59625
diff changeset
    86
  (case SINGLE tac (Goal.init cgoal) of
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    87
    NONE => Fail
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    88
  | SOME st =>
60328
9c94e6a30d29 clarified context;
wenzelm
parents: 59625
diff changeset
    89
      if Thm.no_prems st
9c94e6a30d29 clarified context;
wenzelm
parents: 59625
diff changeset
    90
      then Solved (Goal.finish ctxt st)
9c94e6a30d29 clarified context;
wenzelm
parents: 59625
diff changeset
    91
      else Stuck st)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    92
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    93
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    94
fun dest_binop_list cn (t as (Const (n, _) $ a $ b)) =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    95
  if cn = n then dest_binop_list cn a @ dest_binop_list cn b else [ t ]
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    96
  | dest_binop_list _ t = [ t ]
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    97
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    98
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    99
(* separate two parts in a +-expression:
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   100
   "a + b + c + d + e" --> "(a + b + d) + (c + e)"
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   101
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   102
   Here, + can be any binary operation that is AC.
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   103
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   104
   cn - The name of the binop-constructor (e.g. @{const_name Un})
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   105
   ac - the AC rewrite rules for cn
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   106
   is - the list of indices of the expressions that should become the first part
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   107
        (e.g. [0,1,3] in the above example)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   108
*)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   109
59625
aacdce52b2fc proper context;
wenzelm
parents: 59621
diff changeset
   110
fun regroup_conv ctxt neu cn ac is ct =
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   111
 let
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   112
   val mk = HOLogic.mk_binop cn
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   113
   val t = Thm.term_of ct
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   114
   val xs = dest_binop_list cn t
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   115
   val js = subtract (op =) is (0 upto (length xs) - 1)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   116
   val ty = fastype_of t
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   117
 in
54883
dd04a8b654fc proper context for norm_hhf and derived operations;
wenzelm
parents: 54742
diff changeset
   118
   Goal.prove_internal ctxt []
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   119
     (Thm.cterm_of ctxt
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   120
       (Logic.mk_equals (t,
33611
haftmann
parents: 33099
diff changeset
   121
          if null is
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   122
          then mk (Const (neu, ty), foldr1 mk (map (nth xs) js))
33611
haftmann
parents: 33099
diff changeset
   123
          else if null js
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   124
            then mk (foldr1 mk (map (nth xs) is), Const (neu, ty))
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   125
            else mk (foldr1 mk (map (nth xs) is), foldr1 mk (map (nth xs) js)))))
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54407
diff changeset
   126
     (K (rewrite_goals_tac ctxt ac
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58839
diff changeset
   127
         THEN resolve_tac ctxt [Drule.reflexive_thm] 1))
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   128
 end
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   129
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   130
(* instance for unions *)
59625
aacdce52b2fc proper context;
wenzelm
parents: 59621
diff changeset
   131
fun regroup_union_conv ctxt =
aacdce52b2fc proper context;
wenzelm
parents: 59621
diff changeset
   132
  regroup_conv ctxt @{const_abbrev Set.empty} @{const_name Lattices.sup}
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
   133
    (map (fn t => t RS eq_reflection)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
   134
      (@{thms Un_ac} @ @{thms Un_empty_right} @ @{thms Un_empty_left}))
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   135
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   136
54407
e95831757903 ported part of function package to new 'Ctr_Sugar' abstraction
blanchet
parents: 54406
diff changeset
   137
fun inst_constrs_of ctxt (Type (name, Ts)) =
e95831757903 ported part of function package to new 'Ctr_Sugar' abstraction
blanchet
parents: 54406
diff changeset
   138
    map (Ctr_Sugar.mk_ctr Ts) (#ctrs (the (Ctr_Sugar.ctr_sugar_of ctxt name)))
e95831757903 ported part of function package to new 'Ctr_Sugar' abstraction
blanchet
parents: 54406
diff changeset
   139
  | inst_constrs_of _ _ = raise Match
54406
a2d18fea844a undid copy-paste
blanchet
parents: 45156
diff changeset
   140
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   141
end