src/HOL/Tools/function_package/fundef_lib.ML
author haftmann
Wed, 11 Mar 2009 15:56:50 +0100
changeset 30449 4caf15ae8c26
parent 30304 d8e4cd2ac2a1
child 30763 6976521b4263
permissions -rw-r--r--
tuned
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24584
01e83ffa6c54 fixed title
haftmann
parents: 23819
diff changeset
     1
(*  Title:      HOL/Tools/function_package/fundef_lib.ML
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     2
    Author:     Alexander Krauss, TU Muenchen
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     3
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     4
A package for general recursive function definitions. 
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     5
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
     6
*)
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     7
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20876
diff changeset
     8
structure FundefLib = struct
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20876
diff changeset
     9
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21858
diff changeset
    10
fun map_option f NONE = NONE 
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21858
diff changeset
    11
  | map_option f (SOME x) = SOME (f x);
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21858
diff changeset
    12
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21858
diff changeset
    13
fun fold_option f NONE y = y
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21858
diff changeset
    14
  | fold_option f (SOME x) y = f x y;
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21858
diff changeset
    15
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21858
diff changeset
    16
fun fold_map_option f NONE y = (NONE, y)
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21858
diff changeset
    17
  | fold_map_option f (SOME x) y = apfst SOME (f x y);
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21858
diff changeset
    18
23216
49c78d67b807 added plural (from Pure/library.ML);
wenzelm
parents: 23203
diff changeset
    19
(* Ex: "The variable" ^ plural " is" "s are" vs *)
49c78d67b807 added plural (from Pure/library.ML);
wenzelm
parents: 23203
diff changeset
    20
fun plural sg pl [x] = sg
49c78d67b807 added plural (from Pure/library.ML);
wenzelm
parents: 23203
diff changeset
    21
  | plural sg pl _ = pl
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
    22
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21858
diff changeset
    23
(* lambda-abstracts over an arbitrarily nested tuple
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21858
diff changeset
    24
  ==> hologic.ML? *)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    25
fun tupled_lambda vars t =
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    26
    case vars of
21237
b803f9870e97 untabified
krauss
parents: 21188
diff changeset
    27
      (Free v) => lambda (Free v) t
21188
2aa15b663cd4 minor cleanup
krauss
parents: 21051
diff changeset
    28
    | (Var v) => lambda (Var v) t
2aa15b663cd4 minor cleanup
krauss
parents: 21051
diff changeset
    29
    | (Const ("Pair", Type ("fun", [Ta, Type ("fun", [Tb, _])]))) $ us $ vs =>  
21237
b803f9870e97 untabified
krauss
parents: 21188
diff changeset
    30
      (HOLogic.split_const (Ta,Tb, fastype_of t)) $ (tupled_lambda us (tupled_lambda vs t))
21188
2aa15b663cd4 minor cleanup
krauss
parents: 21051
diff changeset
    31
    | _ => raise Match
2aa15b663cd4 minor cleanup
krauss
parents: 21051
diff changeset
    32
                 
2aa15b663cd4 minor cleanup
krauss
parents: 21051
diff changeset
    33
                 
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    34
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
    35
    let
21237
b803f9870e97 untabified
krauss
parents: 21188
diff changeset
    36
      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
    37
    in
21237
b803f9870e97 untabified
krauss
parents: 21188
diff changeset
    38
      (Free (n, T), body)
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    39
    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    40
  | dest_all _ = raise Match
21188
2aa15b663cd4 minor cleanup
krauss
parents: 21051
diff changeset
    41
                         
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    42
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    43
(* 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
    44
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
    45
    let
21319
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
    46
      val (v,b) = dest_all t
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
    47
      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
    48
    in
21319
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
    49
      (v :: vs, b')
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    50
    end
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    51
  | dest_all_all t = ([],t)
21319
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
    52
                     
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    53
21319
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
    54
(* FIXME: similar to Variable.focus *)
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    55
fun dest_all_all_ctx ctx (Const ("all", _) $ Abs (a as (n,T,b))) =
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    56
    let
20154
c709a29f1363 renamed Variable.rename_wrt to Variable.variant_frees;
wenzelm
parents: 19922
diff changeset
    57
      val [(n', _)] = Variable.variant_frees ctx [] [(n,T)]
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28083
diff changeset
    58
      val (_, ctx') = ProofContext.add_fixes_i [(Binding.name n', SOME T, NoSyn)] ctx
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    59
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    60
      val (n'', body) = Term.dest_abs (n', T, b) 
21858
05f57309170c avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents: 21319
diff changeset
    61
      val _ = (n' = n'') orelse error "dest_all_ctx"
05f57309170c avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents: 21319
diff changeset
    62
      (* Note: We assume that n' does not occur in the body. Otherwise it would be fixed. *)
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    63
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    64
      val (ctx'', vs, bd) = dest_all_all_ctx ctx' body
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    65
    in
21237
b803f9870e97 untabified
krauss
parents: 21188
diff changeset
    66
      (ctx'', (n', T) :: vs, bd)
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    67
    end
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    68
  | dest_all_all_ctx ctx t = 
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    69
    (ctx, [], t)
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    70
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    71
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    72
fun map3 _ [] [] [] = []
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    73
  | 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
    74
  | map3 _ _ _ _ = raise Library.UnequalLengths;
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    75
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
    76
fun map4 _ [] [] [] [] = []
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
    77
  | 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
    78
  | map4 _ _ _ _ _ = raise Library.UnequalLengths;
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
    79
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    80
fun map6 _ [] [] [] [] [] [] = []
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    81
  | 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
    82
  | map6 _ _ _ _ _ _ _ = raise Library.UnequalLengths;
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    83
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
    84
fun map7 _ [] [] [] [] [] [] [] = []
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
    85
  | 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
    86
  | map7 _ _ _ _ _ _ _ _ = raise Library.UnequalLengths;
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
    87
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    88
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    89
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    90
(* forms all "unordered pairs": [1, 2, 3] ==> [(1, 1), (1, 2), (1, 3), (2, 2), (2, 3), (3, 3)] *)
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21858
diff changeset
    91
(* ==> library *)
19922
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    92
fun unordered_pairs [] = []
984ae977f7aa Fixed name clash.
krauss
parents: 19841
diff changeset
    93
  | 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
    94
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19564
diff changeset
    95
27792
e4a4d057749d clean up dead code
krauss
parents: 27790
diff changeset
    96
(* Replaces Frees by name. Works with loose Bounds. *)
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21858
diff changeset
    97
fun replace_frees assoc =
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21858
diff changeset
    98
    map_aterms (fn c as Free (n, _) => the_default c (AList.lookup (op =) assoc n)
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21858
diff changeset
    99
                 | t => t)
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21858
diff changeset
   100
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   101
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   102
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
   103
  | rename_bound n _ = raise Match
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   104
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   105
fun mk_forall_rename (n, v) =
27330
1af2598b5f7d Logic.all/mk_equals/mk_implies;
wenzelm
parents: 24584
diff changeset
   106
    rename_bound n o Logic.all v 
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   107
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   108
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
   109
    let
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   110
      val allthm = forall_intr cv thm
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21858
diff changeset
   111
      val (_ $ abs) = prop_of allthm
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   112
    in
27792
e4a4d057749d clean up dead code
krauss
parents: 27790
diff changeset
   113
      Thm.rename_boundvars abs (Abs (n, dummyT, Term.dummy_pattern dummyT)) allthm
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   114
    end
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   115
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   116
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20154
diff changeset
   117
(* 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
   118
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
   119
    Term.add_frees t []
cf814e36f788 replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents: 21237
diff changeset
   120
    |> 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
   121
    |> rev
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20876
diff changeset
   122
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20876
diff changeset
   123
27790
37b4e65d1dbf FundefLib.try_proof : attempt a proof and see if it works
krauss
parents: 27330
diff changeset
   124
datatype proof_attempt = Solved of thm | Stuck of thm | Fail
37b4e65d1dbf FundefLib.try_proof : attempt a proof and see if it works
krauss
parents: 27330
diff changeset
   125
37b4e65d1dbf FundefLib.try_proof : attempt a proof and see if it works
krauss
parents: 27330
diff changeset
   126
fun try_proof cgoal tac = 
37b4e65d1dbf FundefLib.try_proof : attempt a proof and see if it works
krauss
parents: 27330
diff changeset
   127
    case SINGLE tac (Goal.init cgoal) of
37b4e65d1dbf FundefLib.try_proof : attempt a proof and see if it works
krauss
parents: 27330
diff changeset
   128
      NONE => Fail
37b4e65d1dbf FundefLib.try_proof : attempt a proof and see if it works
krauss
parents: 27330
diff changeset
   129
    | SOME st => if Thm.no_prems st then Solved (Goal.finish st) else Stuck st
37b4e65d1dbf FundefLib.try_proof : attempt a proof and see if it works
krauss
parents: 27330
diff changeset
   130
37b4e65d1dbf FundefLib.try_proof : attempt a proof and see if it works
krauss
parents: 27330
diff changeset
   131
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   132
fun dest_binop_list cn (t as (Const (n, _) $ a $ b)) = 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   133
    if cn = n then dest_binop_list cn a @ dest_binop_list cn b else [ t ]
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   134
  | dest_binop_list _ t = [ t ]
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   135
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   136
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   137
(* separate two parts in a +-expression:
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   138
   "a + b + c + d + e" --> "(a + b + d) + (c + e)"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   139
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   140
   Here, + can be any binary operation that is AC.
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   141
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29302
diff changeset
   142
   cn - The name of the binop-constructor (e.g. @{const_name Un})
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   143
   ac - the AC rewrite rules for cn
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   144
   is - the list of indices of the expressions that should become the first part
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   145
        (e.g. [0,1,3] in the above example)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   146
*)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   147
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   148
fun regroup_conv neu cn ac is ct =
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   149
 let
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   150
   val mk = HOLogic.mk_binop cn
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   151
   val t = term_of ct
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   152
   val xs = dest_binop_list cn t
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   153
   val js = 0 upto (length xs) - 1 \\ is
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   154
   val ty = fastype_of t
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   155
   val thy = theory_of_cterm ct
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   156
 in
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   157
   Goal.prove_internal []
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   158
     (cterm_of thy
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   159
       (Logic.mk_equals (t,
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   160
          if is = []
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   161
          then mk (Const (neu, ty), foldr1 mk (map (nth xs) js))
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   162
          else if js = []
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   163
            then mk (foldr1 mk (map (nth xs) is), Const (neu, ty))
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   164
            else mk (foldr1 mk (map (nth xs) is), foldr1 mk (map (nth xs) js)))))
29302
eb782d1dc07c normalized some ML type/val aliases;
wenzelm
parents: 29125
diff changeset
   165
     (K (rewrite_goals_tac ac
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   166
         THEN rtac Drule.reflexive_thm 1))
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   167
 end
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   168
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   169
(* instance for unions *)
30449
haftmann
parents: 30304
diff changeset
   170
fun regroup_union_conv t = regroup_conv @{const_name Set.empty} @{const_name Un}
haftmann
parents: 30304
diff changeset
   171
  (map (fn t => t RS eq_reflection) (@{thms "Un_ac"} @
haftmann
parents: 30304
diff changeset
   172
                                     @{thms "Un_empty_right"} @
haftmann
parents: 30304
diff changeset
   173
                                     @{thms "Un_empty_left"})) t
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   174
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 28965
diff changeset
   175
21858
05f57309170c avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents: 21319
diff changeset
   176
end