src/HOL/Tools/function_package/fundef_lib.ML
author wenzelm
Wed, 15 Nov 2006 20:50:22 +0100
changeset 21391 a8809f46bd7f
parent 21319 cf814e36f788
child 21858 05f57309170c
permissions -rw-r--r--
replaced NameSpace.append by NameSpace.qualified, which handles empty names as expected;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     1
(*  Title:      HOL/Tools/function_package/lib.ML
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     2
    ID:         $Id$
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     3
    Author:     Alexander Krauss, TU Muenchen
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     4
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     5
A package for general recursive function definitions. 
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     6
Some fairly general functions that should probably go somewhere else... 
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     7
*)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     8
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     9
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20876
diff changeset
    10
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20876
diff changeset
    11
structure FundefLib = struct
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20876
diff changeset
    12
21319
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
    13
fun plural sg pl [x] = sg
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
    14
  | plural sg pl _ = pl
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
    15
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
    16
fun eq_str (s : string, t) = (s = t) (* monomorphic equality on strings *)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
    17
21188
2aa15b663cd4 minor cleanup
krauss
parents: 21051
diff changeset
    18
fun mk_forall v t = all (fastype_of v) $ lambda v t
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    19
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    20
(* Constructs a tupled abstraction from an arbitrarily nested tuple of variables and a term. *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    21
fun tupled_lambda vars t =
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    22
    case vars of
21237
b803f9870e97 untabified
krauss
parents: 21188
diff changeset
    23
      (Free v) => lambda (Free v) t
21188
2aa15b663cd4 minor cleanup
krauss
parents: 21051
diff changeset
    24
    | (Var v) => lambda (Var v) t
2aa15b663cd4 minor cleanup
krauss
parents: 21051
diff changeset
    25
    | (Const ("Pair", Type ("fun", [Ta, Type ("fun", [Tb, _])]))) $ us $ vs =>  
21237
b803f9870e97 untabified
krauss
parents: 21188
diff changeset
    26
      (HOLogic.split_const (Ta,Tb, fastype_of t)) $ (tupled_lambda us (tupled_lambda vs t))
21188
2aa15b663cd4 minor cleanup
krauss
parents: 21051
diff changeset
    27
    | _ => raise Match
2aa15b663cd4 minor cleanup
krauss
parents: 21051
diff changeset
    28
                 
2aa15b663cd4 minor cleanup
krauss
parents: 21051
diff changeset
    29
                 
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    30
fun dest_all (Const ("all", _) $ Abs (a as (_,T,_))) =
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    31
    let
21237
b803f9870e97 untabified
krauss
parents: 21188
diff changeset
    32
      val (n, body) = Term.dest_abs a
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    33
    in
21237
b803f9870e97 untabified
krauss
parents: 21188
diff changeset
    34
      (Free (n, T), body)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    35
    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    36
  | dest_all _ = raise Match
21188
2aa15b663cd4 minor cleanup
krauss
parents: 21051
diff changeset
    37
                         
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    38
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    39
(* Removes all quantifiers from a term, replacing bound variables by frees. *)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    40
fun dest_all_all (t as (Const ("all",_) $ _)) = 
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    41
    let
21319
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
    42
      val (v,b) = dest_all t
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
    43
      val (vs, b') = dest_all_all b
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    44
    in
21319
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
    45
      (v :: vs, b')
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    46
    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    47
  | dest_all_all t = ([],t)
21319
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
    48
                     
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    49
21319
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
    50
(* FIXME: similar to Variable.focus *)
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    51
fun dest_all_all_ctx ctx (Const ("all", _) $ Abs (a as (n,T,b))) =
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    52
    let
20154
c709a29f1363 renamed Variable.rename_wrt to Variable.variant_frees;
wenzelm
parents: 19922
diff changeset
    53
      val [(n', _)] = Variable.variant_frees ctx [] [(n,T)]
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    54
      val (_, ctx') = ProofContext.add_fixes_i [(n', SOME T, NoSyn)] ctx
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    55
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    56
      val (n'', body) = Term.dest_abs (n', T, b) 
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    57
      val _ = assert (n' = n'') "dest_all_ctx" (* Note: We assume that n' does not occur in the body. Otherwise it would be fixed. *)
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    58
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    59
      val (ctx'', vs, bd) = dest_all_all_ctx ctx' body
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    60
    in
21237
b803f9870e97 untabified
krauss
parents: 21188
diff changeset
    61
      (ctx'', (n', T) :: vs, bd)
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    62
    end
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    63
  | dest_all_all_ctx ctx t = 
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    64
    (ctx, [], t)
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    65
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    66
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    67
fun implies_elim_swp a b = implies_elim b a
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    68
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    69
fun map3 _ [] [] [] = []
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    70
  | map3 f (x :: xs) (y :: ys) (z :: zs) = f x y z :: map3 f xs ys zs
19841
f2fa72c13186 avoid unqualified exception;
wenzelm
parents: 19770
diff changeset
    71
  | map3 _ _ _ _ = raise Library.UnequalLengths;
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    72
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
    73
fun map4 _ [] [] [] [] = []
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
    74
  | map4 f (x :: xs) (y :: ys) (z :: zs) (u :: us) = f x y z u :: map4 f xs ys zs us
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
    75
  | map4 _ _ _ _ _ = raise Library.UnequalLengths;
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
    76
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    77
fun map6 _ [] [] [] [] [] [] = []
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    78
  | map6 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (w :: ws) = f x y z u v w :: map6 f xs ys zs us vs ws
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    79
  | map6 _ _ _ _ _ _ _ = raise Library.UnequalLengths;
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    80
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
    81
fun map7 _ [] [] [] [] [] [] [] = []
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
    82
  | map7 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (w :: ws) (b :: bs) = f x y z u v w b :: map7 f xs ys zs us vs ws bs
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
    83
  | map7 _ _ _ _ _ _ _ _ = raise Library.UnequalLengths;
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
    84
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    85
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    86
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    87
(* forms all "unordered pairs": [1, 2, 3] ==> [(1, 1), (1, 2), (1, 3), (2, 2), (2, 3), (3, 3)] *)
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    88
fun unordered_pairs [] = []
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    89
  | unordered_pairs (x::xs) = map (pair x) (x::xs) @ unordered_pairs xs
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    90
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19564
diff changeset
    91
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
    92
(* Replaces Frees by name. Probably quicker than Pattern.rewrite_terms, and works with loose Bounds. *)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
    93
fun replace_frees assoc =
21319
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
    94
    map_aterms (fn c as Free (n, _) => the_default c (AList.lookup (op =) assoc n)
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
    95
                 | t => t)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
    96
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
    97
21319
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
    98
fun forall_aterms P (t $ u) = forall_aterms P t andalso forall_aterms P u
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
    99
  | forall_aterms P (Abs (_, _, t)) = forall_aterms P t
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
   100
  | forall_aterms P a = P a
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
   101
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
   102
fun exists_aterm P = not o forall_aterms (not o P)
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
   103
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
   104
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
   105
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   106
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   107
fun rename_bound n (Q $ Abs(_, T, b)) = (Q $ Abs(n, T, b))
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   108
  | rename_bound n _ = raise Match
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   109
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   110
fun mk_forall_rename (n, v) =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   111
    rename_bound n o mk_forall v 
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   112
21319
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
   113
val dummy_term = Free ("", dummyT)
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
   114
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   115
fun forall_intr_rename (n, cv) thm =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   116
    let
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   117
      val allthm = forall_intr cv thm
21319
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
   118
      val (_, abs) = Logic.dest_all (prop_of allthm)
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   119
    in
21319
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
   120
      Thm.rename_boundvars abs (Abs (n, dummyT, dummy_term)) allthm
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   121
    end
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   122
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   123
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   124
(* Returns the frees in a term in canonical order, excluding the fixes from the context *)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   125
fun frees_in_term ctxt t =
21319
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
   126
    Term.add_frees t []
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
   127
    |> filter_out (Variable.is_fixed ctxt o fst)
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
   128
    |> rev
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
   129
(*
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   130
    rev (fold_aterms (fn  Free (v as (x, _)) =>
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   131
                          if Variable.is_fixed ctxt x then I else insert (fn ((x, _),(y, _)) => x = y) v | _ => I) t [])
21319
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
   132
*)
20851
bf80cb83f8be Improved error reporting
krauss
parents: 20523
diff changeset
   133
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20876
diff changeset
   134
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20876
diff changeset
   135
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20876
diff changeset
   136
end