src/HOL/Tools/Function/function_lib.ML
author wenzelm
Sat, 27 Feb 2010 20:57:08 +0100
changeset 35402 115a5a95710a
parent 34232 36a2a3029fd3
child 36945 9bec62c10714
permissions -rw-r--r--
clarified @{const_name} vs. @{const_abbrev};
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     1
(*  Title:      HOL/Tools/Function/fundef_lib.ML
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
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
     4
A package for general recursive function definitions.
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
     5
Some fairly general functions that should probably go somewhere else...
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
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     8
structure Function_Lib =
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     9
struct
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    10
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    11
fun map_option f NONE = NONE
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    12
  | map_option f (SOME x) = SOME (f x);
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    13
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    14
fun fold_option f NONE y = y
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    15
  | fold_option f (SOME x) y = f x y;
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    16
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    17
(* Ex: "The variable" ^ plural " is" "s are" vs *)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    18
fun plural sg pl [x] = sg
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    19
  | plural sg pl _ = pl
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    20
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    21
(* lambda-abstracts over an arbitrarily nested tuple
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    22
  ==> hologic.ML? *)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    23
fun tupled_lambda vars t =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    24
  case vars of
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    25
    (Free v) => lambda (Free v) t
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    26
  | (Var v) => lambda (Var v) t
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    27
  | (Const ("Pair", Type ("fun", [Ta, Type ("fun", [Tb, _])]))) $ us $ vs =>
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    28
      (HOLogic.split_const (Ta,Tb, fastype_of t)) $ (tupled_lambda us (tupled_lambda vs t))
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    29
  | _ => raise Match
33611
haftmann
parents: 33099
diff changeset
    30
haftmann
parents: 33099
diff changeset
    31
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    32
fun dest_all (Const ("all", _) $ Abs (a as (_,T,_))) =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    33
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    34
    val (n, body) = Term.dest_abs a
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    35
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    36
    (Free (n, T), body)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    37
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    38
  | dest_all _ = raise Match
33611
haftmann
parents: 33099
diff changeset
    39
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    40
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    41
(* Removes all quantifiers from a term, replacing bound variables by frees. *)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    42
fun dest_all_all (t as (Const ("all",_) $ _)) =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    43
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    44
    val (v,b) = dest_all t
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    45
    val (vs, b') = dest_all_all b
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    46
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    47
    (v :: vs, b')
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    48
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    49
  | dest_all_all t = ([],t)
33611
haftmann
parents: 33099
diff changeset
    50
33099
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
(* FIXME: similar to Variable.focus *)
33855
cd8acf137c9c eliminated dead code and some unused bindings, reported by polyml
krauss
parents: 33611
diff changeset
    53
fun dest_all_all_ctx ctx (Const ("all", _) $ Abs (n,T,b)) =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    54
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    55
    val [(n', _)] = Variable.variant_frees ctx [] [(n,T)]
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    56
    val (_, ctx') = ProofContext.add_fixes [(Binding.name n', SOME T, NoSyn)] ctx
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    57
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    58
    val (n'', body) = Term.dest_abs (n', T, b)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    59
    val _ = (n' = n'') orelse error "dest_all_ctx"
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    60
      (* Note: We assume that n' does not occur in the body. Otherwise it would be fixed. *)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    61
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    62
    val (ctx'', vs, bd) = dest_all_all_ctx ctx' body
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    63
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    64
    (ctx'', (n', T) :: vs, bd)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    65
  end
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    66
  | dest_all_all_ctx ctx t =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    67
  (ctx, [], t)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    68
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    69
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    70
fun map3 _ [] [] [] = []
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    71
  | map3 f (x :: xs) (y :: ys) (z :: zs) = f x y z :: map3 f xs ys zs
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    72
  | map3 _ _ _ _ = raise Library.UnequalLengths;
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    73
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    74
fun map4 _ [] [] [] [] = []
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    75
  | map4 f (x :: xs) (y :: ys) (z :: zs) (u :: us) = f x y z u :: map4 f xs ys zs us
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    76
  | map4 _ _ _ _ _ = raise Library.UnequalLengths;
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    77
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    78
fun map6 _ [] [] [] [] [] [] = []
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    79
  | 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
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    80
  | map6 _ _ _ _ _ _ _ = raise Library.UnequalLengths;
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
fun map7 _ [] [] [] [] [] [] [] = []
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    83
  | 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
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    84
  | map7 _ _ _ _ _ _ _ _ = raise Library.UnequalLengths;
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    85
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    86
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    87
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    88
(* 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
    89
fun unordered_pairs [] = []
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    90
  | 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
    91
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
(* Replaces Frees by name. Works with loose Bounds. *)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    94
fun replace_frees assoc =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    95
  map_aterms (fn c as Free (n, _) => the_default c (AList.lookup (op =) assoc n)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    96
    | t => t)
33099
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
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
    99
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
   100
  | rename_bound n _ = raise Match
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
fun mk_forall_rename (n, v) =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
   103
  rename_bound n o Logic.all v
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   104
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   105
fun forall_intr_rename (n, cv) thm =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
   106
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
   107
    val allthm = forall_intr cv thm
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
   108
    val (_ $ abs) = prop_of allthm
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
   109
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
   110
    Thm.rename_boundvars abs (Abs (n, dummyT, Term.dummy_pattern dummyT)) allthm
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
   111
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   112
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   113
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   114
(* Returns the frees in a term in canonical order, excluding the fixes from the context *)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   115
fun frees_in_term ctxt t =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
   116
  Term.add_frees t []
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
   117
  |> filter_out (Variable.is_fixed ctxt o fst)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
   118
  |> rev
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   119
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   120
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   121
datatype proof_attempt = Solved of thm | Stuck of thm | Fail
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   122
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
   123
fun try_proof cgoal tac =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
   124
  case SINGLE tac (Goal.init cgoal) of
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
   125
    NONE => Fail
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
   126
  | SOME st =>
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
   127
    if Thm.no_prems st
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
   128
    then Solved (Goal.finish (Syntax.init_pretty_global (Thm.theory_of_cterm cgoal)) st)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
   129
    else Stuck st
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   130
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   131
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
   132
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
   133
  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
   134
  | dest_binop_list _ t = [ t ]
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
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   137
(* separate two parts in a +-expression:
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   138
   "a + b + c + d + e" --> "(a + b + d) + (c + e)"
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   139
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   140
   Here, + can be any binary operation that is AC.
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   141
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   142
   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
   143
   ac - the AC rewrite rules for cn
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   144
   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
   145
        (e.g. [0,1,3] in the above example)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   146
*)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   147
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   148
fun regroup_conv neu cn ac is ct =
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   149
 let
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   150
   val mk = HOLogic.mk_binop cn
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   151
   val t = term_of ct
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   152
   val xs = dest_binop_list cn t
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   153
   val js = subtract (op =) is (0 upto (length xs) - 1)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   154
   val ty = fastype_of t
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   155
   val thy = theory_of_cterm ct
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   156
 in
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   157
   Goal.prove_internal []
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   158
     (cterm_of thy
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   159
       (Logic.mk_equals (t,
33611
haftmann
parents: 33099
diff changeset
   160
          if null is
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   161
          then mk (Const (neu, ty), foldr1 mk (map (nth xs) js))
33611
haftmann
parents: 33099
diff changeset
   162
          else if null js
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   163
            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
   164
            else mk (foldr1 mk (map (nth xs) is), foldr1 mk (map (nth xs) js)))))
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   165
     (K (rewrite_goals_tac ac
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   166
         THEN rtac Drule.reflexive_thm 1))
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   167
 end
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   168
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   169
(* instance for unions *)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
   170
val regroup_union_conv =
35402
115a5a95710a clarified @{const_name} vs. @{const_abbrev};
wenzelm
parents: 34232
diff changeset
   171
  regroup_conv @{const_abbrev Set.empty} @{const_name Lattices.sup}
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
   172
    (map (fn t => t RS eq_reflection)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 33855
diff changeset
   173
      (@{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
   174
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   175
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   176
end