src/HOL/Product_Type.thy
author wenzelm
Fri, 21 Oct 2005 18:14:34 +0200
changeset 17956 369e2af8ee45
parent 17875 d81094515061
child 18013 3f5d0acdfdba
permissions -rw-r--r--
Goal.prove;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/Product_Type.thy
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     2
    ID:         $Id$
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
11777
wenzelm
parents: 11602
diff changeset
     5
*)
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     6
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
     7
header {* Cartesian products *}
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     8
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14952
diff changeset
     9
theory Product_Type
15140
322485b816ac import -> imports
nipkow
parents: 15131
diff changeset
    10
imports Fun
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 16121
diff changeset
    11
uses ("Tools/split_rule.ML")
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14952
diff changeset
    12
begin
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    13
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    14
subsection {* Unit *}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    15
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    16
typedef unit = "{True}"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    17
proof
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    18
  show "True : ?unit" by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    19
qed
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    20
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    21
constdefs
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    22
  Unity :: unit    ("'(')")
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    23
  "() == Abs_unit True"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    24
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    25
lemma unit_eq: "u = ()"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    26
  by (induct u) (simp add: unit_def Unity_def)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    27
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    28
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    29
  Simplification procedure for @{thm [source] unit_eq}.  Cannot use
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    30
  this rule directly --- it loops!
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    31
*}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    32
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    33
ML_setup {*
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12338
diff changeset
    34
  val unit_eq_proc =
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12338
diff changeset
    35
    let val unit_meta_eq = mk_meta_eq (thm "unit_eq") in
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12338
diff changeset
    36
      Simplifier.simproc (Theory.sign_of (the_context ())) "unit_eq" ["x::unit"]
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
    37
      (fn _ => fn _ => fn t => if HOLogic.is_unit t then NONE else SOME unit_meta_eq)
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12338
diff changeset
    38
    end;
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    39
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    40
  Addsimprocs [unit_eq_proc];
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    41
*}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    42
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    43
lemma unit_all_eq1: "(!!x::unit. PROP P x) == PROP P ()"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    44
  by simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    45
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    46
lemma unit_all_eq2: "(!!x::unit. PROP P) == PROP P"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    47
  by (rule triv_forall_equality)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    48
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    49
lemma unit_induct [induct type: unit]: "P () ==> P x"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    50
  by simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    51
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    52
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    53
  This rewrite counters the effect of @{text unit_eq_proc} on @{term
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    54
  [source] "%u::unit. f u"}, replacing it by @{term [source]
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    55
  f} rather than by @{term [source] "%u. f ()"}.
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    56
*}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    57
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    58
lemma unit_abs_eta_conv [simp]: "(%u::unit. f ()) = f"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    59
  by (rule ext) simp
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    60
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    61
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    62
subsection {* Pairs *}
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    63
11777
wenzelm
parents: 11602
diff changeset
    64
subsubsection {* Type definition *}
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    65
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    66
constdefs
11025
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
    67
  Pair_Rep :: "['a, 'b] => ['a, 'b] => bool"
11032
83f723e86dac added hidden internal_split constant;
wenzelm
parents: 11025
diff changeset
    68
  "Pair_Rep == (%a b. %x y. x=a & y=b)"
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    69
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    70
global
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    71
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    72
typedef (Prod)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    73
  ('a, 'b) "*"    (infixr 20)
11032
83f723e86dac added hidden internal_split constant;
wenzelm
parents: 11025
diff changeset
    74
    = "{f. EX a b. f = Pair_Rep (a::'a) (b::'b)}"
11025
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
    75
proof
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
    76
  fix a b show "Pair_Rep a b : ?Prod"
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
    77
    by blast
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
    78
qed
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    79
12114
a8e860c86252 eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents: 11966
diff changeset
    80
syntax (xsymbols)
15422
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
    81
  "*"      :: "[type, type] => type"         ("(_ \<times>/ _)" [21, 20] 20)
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    82
syntax (HTML output)
15422
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
    83
  "*"      :: "[type, type] => type"         ("(_ \<times>/ _)" [21, 20] 20)
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    84
11777
wenzelm
parents: 11602
diff changeset
    85
local
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    86
11777
wenzelm
parents: 11602
diff changeset
    87
wenzelm
parents: 11602
diff changeset
    88
subsubsection {* Abstract constants and syntax *}
wenzelm
parents: 11602
diff changeset
    89
wenzelm
parents: 11602
diff changeset
    90
global
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    91
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    92
consts
11025
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
    93
  fst      :: "'a * 'b => 'a"
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
    94
  snd      :: "'a * 'b => 'b"
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
    95
  split    :: "[['a, 'b] => 'c, 'a * 'b] => 'c"
14189
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
    96
  curry    :: "['a * 'b => 'c, 'a, 'b] => 'c"
11025
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
    97
  prod_fun :: "['a => 'b, 'c => 'd, 'a * 'c] => 'b * 'd"
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
    98
  Pair     :: "['a, 'b] => 'a * 'b"
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
    99
  Sigma    :: "['a set, 'a => 'b set] => ('a * 'b) set"
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   100
11777
wenzelm
parents: 11602
diff changeset
   101
local
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   102
11777
wenzelm
parents: 11602
diff changeset
   103
text {*
wenzelm
parents: 11602
diff changeset
   104
  Patterns -- extends pre-defined type @{typ pttrn} used in
wenzelm
parents: 11602
diff changeset
   105
  abstractions.
wenzelm
parents: 11602
diff changeset
   106
*}
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   107
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   108
nonterminals
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   109
  tuple_args patterns
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   110
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   111
syntax
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   112
  "_tuple"      :: "'a => tuple_args => 'a * 'b"        ("(1'(_,/ _'))")
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   113
  "_tuple_arg"  :: "'a => tuple_args"                   ("_")
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   114
  "_tuple_args" :: "'a => tuple_args => tuple_args"     ("_,/ _")
11025
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   115
  "_pattern"    :: "[pttrn, patterns] => pttrn"         ("'(_,/ _')")
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   116
  ""            :: "pttrn => patterns"                  ("_")
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   117
  "_patterns"   :: "[pttrn, patterns] => patterns"      ("_,/ _")
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   118
  "@Sigma" ::"[pttrn, 'a set, 'b set] => ('a * 'b) set" ("(3SIGMA _:_./ _)" 10)
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   119
  "@Times" ::"['a set,  'a => 'b set] => ('a * 'b) set" (infixr "<*>" 80)
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   120
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   121
translations
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   122
  "(x, y)"       == "Pair x y"
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   123
  "_tuple x (_tuple_args y z)" == "_tuple x (_tuple_arg (_tuple y z))"
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   124
  "%(x,y,zs).b"  == "split(%x (y,zs).b)"
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   125
  "%(x,y).b"     == "split(%x y. b)"
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   126
  "_abs (Pair x y) t" => "%(x,y).t"
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   127
  (* The last rule accommodates tuples in `case C ... (x,y) ... => ...'
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   128
     The (x,y) is parsed as `Pair x y' because it is logic, not pttrn *)
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   129
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   130
  "SIGMA x:A. B" => "Sigma A (%x. B)"
17782
b3846df9d643 replaced _K by dummy abstraction;
wenzelm
parents: 17085
diff changeset
   131
  "A <*> B"      => "Sigma A (%_. B)"
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   132
14359
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   133
(* reconstructs pattern from (nested) splits, avoiding eta-contraction of body*)
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   134
(* works best with enclosing "let", if "let" does not avoid eta-contraction   *)
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   135
print_translation {*
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   136
let fun split_tr' [Abs (x,T,t as (Abs abs))] =
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   137
      (* split (%x y. t) => %(x,y) t *)
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   138
      let val (y,t') = atomic_abs_tr' abs;
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   139
          val (x',t'') = atomic_abs_tr' (x,T,t');
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   140
    
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   141
      in Syntax.const "_abs" $ (Syntax.const "_pattern" $x'$y) $ t'' end
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   142
    | split_tr' [Abs (x,T,(s as Const ("split",_)$t))] =
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   143
       (* split (%x. (split (%y z. t))) => %(x,y,z). t *)
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   144
       let val (Const ("_abs",_)$(Const ("_pattern",_)$y$z)$t') = split_tr' [t];
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   145
           val (x',t'') = atomic_abs_tr' (x,T,t');
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   146
       in Syntax.const "_abs"$ 
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   147
           (Syntax.const "_pattern"$x'$(Syntax.const "_patterns"$y$z))$t'' end
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   148
    | split_tr' [Const ("split",_)$t] =
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   149
       (* split (split (%x y z. t)) => %((x,y),z). t *)   
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   150
       split_tr' [(split_tr' [t])] (* inner split_tr' creates next pattern *)
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   151
    | split_tr' [Const ("_abs",_)$x_y$(Abs abs)] =
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   152
       (* split (%pttrn z. t) => %(pttrn,z). t *)
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   153
       let val (z,t) = atomic_abs_tr' abs;
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   154
       in Syntax.const "_abs" $ (Syntax.const "_pattern" $x_y$z) $ t end
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   155
    | split_tr' _ =  raise Match;
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   156
in [("split", split_tr')]
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   157
end
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   158
*}
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   159
15422
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   160
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   161
(* print "split f" as "\<lambda>(x,y). f x y" and "split (\<lambda>x. f x)" as "\<lambda>(x,y). f x y" *) 
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   162
typed_print_translation {*
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   163
let
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   164
  fun split_guess_names_tr' _ T [Abs (x,_,Abs _)] = raise Match
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   165
    | split_guess_names_tr' _ T  [Abs (x,xT,t)] =
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   166
        (case (head_of t) of
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   167
           Const ("split",_) => raise Match
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   168
         | _ => let 
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   169
                  val (_::yT::_) = binder_types (domain_type T) handle Bind => raise Match;
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   170
                  val (y,t') = atomic_abs_tr' ("y",yT,(incr_boundvars 1 t)$Bound 0); 
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   171
                  val (x',t'') = atomic_abs_tr' (x,xT,t');
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   172
                in Syntax.const "_abs" $ (Syntax.const "_pattern" $x'$y) $ t'' end)
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   173
    | split_guess_names_tr' _ T [t] =
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   174
       (case (head_of t) of
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   175
           Const ("split",_) => raise Match 
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   176
         | _ => let 
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   177
                  val (xT::yT::_) = binder_types (domain_type T) handle Bind => raise Match;
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   178
                  val (y,t') = 
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   179
                        atomic_abs_tr' ("y",yT,(incr_boundvars 2 t)$Bound 1$Bound 0); 
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   180
                  val (x',t'') = atomic_abs_tr' ("x",xT,t');
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   181
                in Syntax.const "_abs" $ (Syntax.const "_pattern" $x'$y) $ t'' end)
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   182
    | split_guess_names_tr' _ _ _ = raise Match;
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   183
in [("split", split_guess_names_tr')]
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   184
end 
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   185
*}
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   186
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   187
text{*Deleted x-symbol and html support using @{text"\<Sigma>"} (Sigma) because of the danger of confusion with Sum.*}
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   188
12114
a8e860c86252 eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents: 11966
diff changeset
   189
syntax (xsymbols)
15422
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   190
  "@Times" :: "['a set,  'a => 'b set] => ('a * 'b) set"  ("_ \<times> _" [81, 80] 80)
11025
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   191
14565
c6dc17aab88a use more symbols in HTML output
kleing
parents: 14359
diff changeset
   192
syntax (HTML output)
15422
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   193
  "@Times" :: "['a set,  'a => 'b set] => ('a * 'b) set"  ("_ \<times> _" [81, 80] 80)
14565
c6dc17aab88a use more symbols in HTML output
kleing
parents: 14359
diff changeset
   194
11032
83f723e86dac added hidden internal_split constant;
wenzelm
parents: 11025
diff changeset
   195
print_translation {* [("Sigma", dependent_tr' ("@Sigma", "@Times"))] *}
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   196
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   197
11777
wenzelm
parents: 11602
diff changeset
   198
subsubsection {* Definitions *}
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   199
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   200
defs
11025
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   201
  Pair_def:     "Pair a b == Abs_Prod(Pair_Rep a b)"
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 11425
diff changeset
   202
  fst_def:      "fst p == THE a. EX b. p = (a, b)"
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 11425
diff changeset
   203
  snd_def:      "snd p == THE b. EX a. p = (a, b)"
11025
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   204
  split_def:    "split == (%c p. c (fst p) (snd p))"
14189
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
   205
  curry_def:    "curry == (%c x y. c (x,y))"
11025
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   206
  prod_fun_def: "prod_fun f g == split(%x y.(f(x), g(y)))"
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   207
  Sigma_def:    "Sigma A B == UN x:A. UN y:B(x). {(x, y)}"
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   208
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   209
11966
wenzelm
parents: 11838
diff changeset
   210
subsubsection {* Lemmas and proof tool setup *}
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   211
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   212
lemma ProdI: "Pair_Rep a b : Prod"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   213
  by (unfold Prod_def) blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   214
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   215
lemma Pair_Rep_inject: "Pair_Rep a b = Pair_Rep a' b' ==> a = a' & b = b'"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   216
  apply (unfold Pair_Rep_def)
14208
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   217
  apply (drule fun_cong [THEN fun_cong], blast)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   218
  done
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   219
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   220
lemma inj_on_Abs_Prod: "inj_on Abs_Prod Prod"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   221
  apply (rule inj_on_inverseI)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   222
  apply (erule Abs_Prod_inverse)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   223
  done
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   224
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   225
lemma Pair_inject:
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   226
  "(a, b) = (a', b') ==> (a = a' ==> b = b' ==> R) ==> R"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   227
proof -
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   228
  case rule_context [unfolded Pair_def]
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   229
  show ?thesis
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   230
    apply (rule inj_on_Abs_Prod [THEN inj_onD, THEN Pair_Rep_inject, THEN conjE])
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   231
    apply (rule rule_context ProdI)+
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   232
    .
11025
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   233
qed
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   234
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   235
lemma Pair_eq [iff]: "((a, b) = (a', b')) = (a = a' & b = b')"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   236
  by (blast elim!: Pair_inject)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   237
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   238
lemma fst_conv [simp]: "fst (a, b) = a"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   239
  by (unfold fst_def) blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   240
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   241
lemma snd_conv [simp]: "snd (a, b) = b"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   242
  by (unfold snd_def) blast
11025
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   243
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   244
lemma fst_eqD: "fst (x, y) = a ==> x = a"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   245
  by simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   246
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   247
lemma snd_eqD: "snd (x, y) = a ==> y = a"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   248
  by simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   249
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   250
lemma PairE_lemma: "EX x y. p = (x, y)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   251
  apply (unfold Pair_def)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   252
  apply (rule Rep_Prod [unfolded Prod_def, THEN CollectE])
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   253
  apply (erule exE, erule exE, rule exI, rule exI)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   254
  apply (rule Rep_Prod_inverse [symmetric, THEN trans])
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   255
  apply (erule arg_cong)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   256
  done
11032
83f723e86dac added hidden internal_split constant;
wenzelm
parents: 11025
diff changeset
   257
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   258
lemma PairE [cases type: *]: "(!!x y. p = (x, y) ==> Q) ==> Q"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   259
  by (insert PairE_lemma [of p]) blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   260
16121
wenzelm
parents: 15570
diff changeset
   261
ML {*
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   262
  local val PairE = thm "PairE" in
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   263
    fun pair_tac s =
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   264
      EVERY' [res_inst_tac [("p", s)] PairE, hyp_subst_tac, K prune_params_tac];
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   265
  end;
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   266
*}
11032
83f723e86dac added hidden internal_split constant;
wenzelm
parents: 11025
diff changeset
   267
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   268
lemma surjective_pairing: "p = (fst p, snd p)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   269
  -- {* Do not add as rewrite rule: invalidates some proofs in IMP *}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   270
  by (cases p) simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   271
17085
5b57f995a179 more simprules now have names
paulson
parents: 17021
diff changeset
   272
lemmas pair_collapse = surjective_pairing [symmetric]
5b57f995a179 more simprules now have names
paulson
parents: 17021
diff changeset
   273
declare pair_collapse [simp]
11025
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   274
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   275
lemma surj_pair [simp]: "EX x y. z = (x, y)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   276
  apply (rule exI)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   277
  apply (rule exI)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   278
  apply (rule surjective_pairing)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   279
  done
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   280
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   281
lemma split_paired_all: "(!!x. PROP P x) == (!!a b. PROP P (a, b))"
11820
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   282
proof
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   283
  fix a b
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   284
  assume "!!x. PROP P x"
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   285
  thus "PROP P (a, b)" .
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   286
next
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   287
  fix x
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   288
  assume "!!a b. PROP P (a, b)"
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   289
  hence "PROP P (fst x, snd x)" .
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   290
  thus "PROP P x" by simp
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   291
qed
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   292
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   293
lemmas split_tupled_all = split_paired_all unit_all_eq2
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   294
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   295
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   296
  The rule @{thm [source] split_paired_all} does not work with the
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   297
  Simplifier because it also affects premises in congrence rules,
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   298
  where this can lead to premises of the form @{text "!!a b. ... =
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   299
  ?P(a, b)"} which cannot be solved by reflexivity.
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   300
*}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   301
16121
wenzelm
parents: 15570
diff changeset
   302
ML_setup {*
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   303
  (* replace parameters of product type by individual component parameters *)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   304
  val safe_full_simp_tac = generic_simp_tac true (true, false, false);
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   305
  local (* filtering with exists_paired_all is an essential optimization *)
16121
wenzelm
parents: 15570
diff changeset
   306
    fun exists_paired_all (Const ("all", _) $ Abs (_, T, t)) =
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   307
          can HOLogic.dest_prodT T orelse exists_paired_all t
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   308
      | exists_paired_all (t $ u) = exists_paired_all t orelse exists_paired_all u
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   309
      | exists_paired_all (Abs (_, _, t)) = exists_paired_all t
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   310
      | exists_paired_all _ = false;
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   311
    val ss = HOL_basic_ss
16121
wenzelm
parents: 15570
diff changeset
   312
      addsimps [thm "split_paired_all", thm "unit_all_eq2", thm "unit_abs_eta_conv"]
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   313
      addsimprocs [unit_eq_proc];
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   314
  in
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   315
    val split_all_tac = SUBGOAL (fn (t, i) =>
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   316
      if exists_paired_all t then safe_full_simp_tac ss i else no_tac);
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   317
    val unsafe_split_all_tac = SUBGOAL (fn (t, i) =>
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   318
      if exists_paired_all t then full_simp_tac ss i else no_tac);
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   319
    fun split_all th =
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   320
   if exists_paired_all (#prop (Thm.rep_thm th)) then full_simplify ss th else th;
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   321
  end;
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   322
17875
d81094515061 change_claset/simpset;
wenzelm
parents: 17782
diff changeset
   323
change_claset (fn cs => cs addSbefore ("split_all_tac", split_all_tac));
16121
wenzelm
parents: 15570
diff changeset
   324
*}
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   325
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   326
lemma split_paired_All [simp]: "(ALL x. P x) = (ALL a b. P (a, b))"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   327
  -- {* @{text "[iff]"} is not a good idea because it makes @{text blast} loop *}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   328
  by fast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   329
14189
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
   330
lemma curry_split [simp]: "curry (split f) = f"
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
   331
  by (simp add: curry_def split_def)
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
   332
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
   333
lemma split_curry [simp]: "split (curry f) = f"
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
   334
  by (simp add: curry_def split_def)
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
   335
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
   336
lemma curryI [intro!]: "f (a,b) ==> curry f a b"
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
   337
  by (simp add: curry_def)
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
   338
14190
609c072edf90 Fixed blunder in the setup of the classical reasoner wrt. the constant
skalberg
parents: 14189
diff changeset
   339
lemma curryD [dest!]: "curry f a b ==> f (a,b)"
14189
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
   340
  by (simp add: curry_def)
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
   341
14190
609c072edf90 Fixed blunder in the setup of the classical reasoner wrt. the constant
skalberg
parents: 14189
diff changeset
   342
lemma curryE: "[| curry f a b ; f (a,b) ==> Q |] ==> Q"
14189
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
   343
  by (simp add: curry_def)
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
   344
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
   345
lemma curry_conv [simp]: "curry f a b = f (a,b)"
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
   346
  by (simp add: curry_def)
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
   347
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   348
lemma prod_induct [induct type: *]: "!!x. (!!a b. P (a, b)) ==> P x"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   349
  by fast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   350
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   351
lemma split_paired_Ex [simp]: "(EX x. P x) = (EX a b. P (a, b))"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   352
  by fast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   353
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   354
lemma split_conv [simp]: "split c (a, b) = c a b"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   355
  by (simp add: split_def)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   356
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   357
lemmas split = split_conv  -- {* for backwards compatibility *}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   358
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   359
lemmas splitI = split_conv [THEN iffD2, standard]
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   360
lemmas splitD = split_conv [THEN iffD1, standard]
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   361
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   362
lemma split_Pair_apply: "split (%x y. f (x, y)) = f"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   363
  -- {* Subsumes the old @{text split_Pair} when @{term f} is the identity function. *}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   364
  apply (rule ext)
14208
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   365
  apply (tactic {* pair_tac "x" 1 *}, simp)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   366
  done
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   367
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   368
lemma split_paired_The: "(THE x. P x) = (THE (a, b). P (a, b))"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   369
  -- {* Can't be added to simpset: loops! *}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   370
  by (simp add: split_Pair_apply)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   371
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   372
lemma The_split: "The (split P) = (THE xy. P (fst xy) (snd xy))"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   373
  by (simp add: split_def)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   374
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   375
lemma Pair_fst_snd_eq: "!!s t. (s = t) = (fst s = fst t & snd s = snd t)"
14208
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   376
by (simp only: split_tupled_all, simp)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   377
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   378
lemma prod_eqI [intro?]: "fst p = fst q ==> snd p = snd q ==> p = q"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   379
  by (simp add: Pair_fst_snd_eq)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   380
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   381
lemma split_weak_cong: "p = q ==> split c p = split c q"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   382
  -- {* Prevents simplification of @{term c}: much faster *}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   383
  by (erule arg_cong)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   384
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   385
lemma split_eta: "(%(x, y). f (x, y)) = f"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   386
  apply (rule ext)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   387
  apply (simp only: split_tupled_all)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   388
  apply (rule split_conv)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   389
  done
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   390
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   391
lemma cond_split_eta: "(!!x y. f x y = g (x, y)) ==> (%(x, y). f x y) = g"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   392
  by (simp add: split_eta)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   393
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   394
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   395
  Simplification procedure for @{thm [source] cond_split_eta}.  Using
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   396
  @{thm [source] split_eta} as a rewrite rule is not general enough,
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   397
  and using @{thm [source] cond_split_eta} directly would render some
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   398
  existing proofs very inefficient; similarly for @{text
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   399
  split_beta}. *}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   400
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   401
ML_setup {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   402
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   403
local
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   404
  val cond_split_eta = thm "cond_split_eta";
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   405
  fun  Pair_pat k 0 (Bound m) = (m = k)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   406
  |    Pair_pat k i (Const ("Pair",  _) $ Bound m $ t) = i > 0 andalso
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   407
                        m = k+i andalso Pair_pat k (i-1) t
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   408
  |    Pair_pat _ _ _ = false;
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   409
  fun no_args k i (Abs (_, _, t)) = no_args (k+1) i t
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   410
  |   no_args k i (t $ u) = no_args k i t andalso no_args k i u
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   411
  |   no_args k i (Bound m) = m < k orelse m > k+i
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   412
  |   no_args _ _ _ = true;
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   413
  fun split_pat tp i (Abs  (_,_,t)) = if tp 0 i t then SOME (i,t) else NONE
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   414
  |   split_pat tp i (Const ("split", _) $ Abs (_, _, t)) = split_pat tp (i+1) t
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   415
  |   split_pat tp i _ = NONE;
17956
369e2af8ee45 Goal.prove;
wenzelm
parents: 17875
diff changeset
   416
  fun metaeq thy ss lhs rhs = mk_meta_eq (Goal.prove thy [] []
13480
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13462
diff changeset
   417
        (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs,rhs)))
17875
d81094515061 change_claset/simpset;
wenzelm
parents: 17782
diff changeset
   418
        (K (simp_tac (Simplifier.inherit_context ss HOL_basic_ss addsimps [cond_split_eta]) 1)));
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   419
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   420
  fun beta_term_pat k i (Abs (_, _, t)) = beta_term_pat (k+1) i t
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   421
  |   beta_term_pat k i (t $ u) = Pair_pat k i (t $ u) orelse
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   422
                        (beta_term_pat k i t andalso beta_term_pat k i u)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   423
  |   beta_term_pat k i t = no_args k i t;
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   424
  fun  eta_term_pat k i (f $ arg) = no_args k i f andalso Pair_pat k i arg
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   425
  |    eta_term_pat _ _ _ = false;
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   426
  fun subst arg k i (Abs (x, T, t)) = Abs (x, T, subst arg (k+1) i t)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   427
  |   subst arg k i (t $ u) = if Pair_pat k i (t $ u) then incr_boundvars k arg
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   428
                              else (subst arg k i t $ subst arg k i u)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   429
  |   subst arg k i t = t;
17002
fb9261990ffe simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 16770
diff changeset
   430
  fun beta_proc thy ss (s as Const ("split", _) $ Abs (_, _, t) $ arg) =
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   431
        (case split_pat beta_term_pat 1 t of
17002
fb9261990ffe simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 16770
diff changeset
   432
        SOME (i,f) => SOME (metaeq thy ss s (subst arg 0 i f))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   433
        | NONE => NONE)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   434
  |   beta_proc _ _ _ = NONE;
17002
fb9261990ffe simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 16770
diff changeset
   435
  fun eta_proc thy ss (s as Const ("split", _) $ Abs (_, _, t)) =
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   436
        (case split_pat eta_term_pat 1 t of
17002
fb9261990ffe simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 16770
diff changeset
   437
          SOME (_,ft) => SOME (metaeq thy ss s (let val (f $ arg) = ft in f end))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   438
        | NONE => NONE)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   439
  |   eta_proc _ _ _ = NONE;
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   440
in
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12338
diff changeset
   441
  val split_beta_proc = Simplifier.simproc (Theory.sign_of (the_context ()))
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12338
diff changeset
   442
    "split_beta" ["split f z"] beta_proc;
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12338
diff changeset
   443
  val split_eta_proc = Simplifier.simproc (Theory.sign_of (the_context ()))
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12338
diff changeset
   444
    "split_eta" ["split f"] eta_proc;
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   445
end;
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   446
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   447
Addsimprocs [split_beta_proc, split_eta_proc];
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   448
*}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   449
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   450
lemma split_beta: "(%(x, y). P x y) z = P (fst z) (snd z)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   451
  by (subst surjective_pairing, rule split_conv)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   452
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   453
lemma split_split: "R (split c p) = (ALL x y. p = (x, y) --> R (c x y))"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   454
  -- {* For use with @{text split} and the Simplifier. *}
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15422
diff changeset
   455
  by (insert surj_pair [of p], clarify, simp)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   456
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   457
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   458
  @{thm [source] split_split} could be declared as @{text "[split]"}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   459
  done after the Splitter has been speeded up significantly;
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   460
  precompute the constants involved and don't do anything unless the
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   461
  current goal contains one of those constants.
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   462
*}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   463
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   464
lemma split_split_asm: "R (split c p) = (~(EX x y. p = (x, y) & (~R (c x y))))"
14208
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   465
by (subst split_split, simp)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   466
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   467
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   468
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   469
  \medskip @{term split} used as a logical connective or set former.
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   470
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   471
  \medskip These rules are for use with @{text blast}; could instead
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   472
  call @{text simp} using @{thm [source] split} as rewrite. *}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   473
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   474
lemma splitI2: "!!p. [| !!a b. p = (a, b) ==> c a b |] ==> split c p"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   475
  apply (simp only: split_tupled_all)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   476
  apply (simp (no_asm_simp))
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   477
  done
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   478
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   479
lemma splitI2': "!!p. [| !!a b. (a, b) = p ==> c a b x |] ==> split c p x"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   480
  apply (simp only: split_tupled_all)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   481
  apply (simp (no_asm_simp))
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   482
  done
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   483
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   484
lemma splitE: "split c p ==> (!!x y. p = (x, y) ==> c x y ==> Q) ==> Q"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   485
  by (induct p) (auto simp add: split_def)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   486
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   487
lemma splitE': "split c p z ==> (!!x y. p = (x, y) ==> c x y z ==> Q) ==> Q"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   488
  by (induct p) (auto simp add: split_def)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   489
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   490
lemma splitE2:
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   491
  "[| Q (split P z);  !!x y. [|z = (x, y); Q (P x y)|] ==> R |] ==> R"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   492
proof -
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   493
  assume q: "Q (split P z)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   494
  assume r: "!!x y. [|z = (x, y); Q (P x y)|] ==> R"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   495
  show R
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   496
    apply (rule r surjective_pairing)+
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   497
    apply (rule split_beta [THEN subst], rule q)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   498
    done
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   499
qed
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   500
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   501
lemma splitD': "split R (a,b) c ==> R a b c"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   502
  by simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   503
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   504
lemma mem_splitI: "z: c a b ==> z: split c (a, b)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   505
  by simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   506
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   507
lemma mem_splitI2: "!!p. [| !!a b. p = (a, b) ==> z: c a b |] ==> z: split c p"
14208
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   508
by (simp only: split_tupled_all, simp)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   509
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   510
lemma mem_splitE: "[| z: split c p; !!x y. [| p = (x,y); z: c x y |] ==> Q |] ==> Q"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   511
proof -
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   512
  case rule_context [unfolded split_def]
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   513
  show ?thesis by (rule rule_context surjective_pairing)+
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   514
qed
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   515
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   516
declare mem_splitI2 [intro!] mem_splitI [intro!] splitI2' [intro!] splitI2 [intro!] splitI [intro!]
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   517
declare mem_splitE [elim!] splitE' [elim!] splitE [elim!]
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   518
16121
wenzelm
parents: 15570
diff changeset
   519
ML_setup {*
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   520
local (* filtering with exists_p_split is an essential optimization *)
16121
wenzelm
parents: 15570
diff changeset
   521
  fun exists_p_split (Const ("split",_) $ _ $ (Const ("Pair",_)$_$_)) = true
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   522
    | exists_p_split (t $ u) = exists_p_split t orelse exists_p_split u
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   523
    | exists_p_split (Abs (_, _, t)) = exists_p_split t
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   524
    | exists_p_split _ = false;
16121
wenzelm
parents: 15570
diff changeset
   525
  val ss = HOL_basic_ss addsimps [thm "split_conv"];
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   526
in
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   527
val split_conv_tac = SUBGOAL (fn (t, i) =>
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   528
    if exists_p_split t then safe_full_simp_tac ss i else no_tac);
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   529
end;
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   530
(* This prevents applications of splitE for already splitted arguments leading
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   531
   to quite time-consuming computations (in particular for nested tuples) *)
17875
d81094515061 change_claset/simpset;
wenzelm
parents: 17782
diff changeset
   532
change_claset (fn cs => cs addSbefore ("split_conv_tac", split_conv_tac));
16121
wenzelm
parents: 15570
diff changeset
   533
*}
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   534
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   535
lemma split_eta_SetCompr [simp]: "(%u. EX x y. u = (x, y) & P (x, y)) = P"
14208
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   536
by (rule ext, fast)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   537
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   538
lemma split_eta_SetCompr2 [simp]: "(%u. EX x y. u = (x, y) & P x y) = split P"
14208
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   539
by (rule ext, fast)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   540
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   541
lemma split_part [simp]: "(%(a,b). P & Q a b) = (%ab. P & split Q ab)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   542
  -- {* Allows simplifications of nested splits in case of independent predicates. *}
14208
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   543
  apply (rule ext, blast)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   544
  done
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   545
14337
e13731554e50 undid split_comp_eq[simp] because it leads to nontermination together with split_def!
nipkow
parents: 14208
diff changeset
   546
(* Do NOT make this a simp rule as it
e13731554e50 undid split_comp_eq[simp] because it leads to nontermination together with split_def!
nipkow
parents: 14208
diff changeset
   547
   a) only helps in special situations
e13731554e50 undid split_comp_eq[simp] because it leads to nontermination together with split_def!
nipkow
parents: 14208
diff changeset
   548
   b) can lead to nontermination in the presence of split_def
e13731554e50 undid split_comp_eq[simp] because it leads to nontermination together with split_def!
nipkow
parents: 14208
diff changeset
   549
*)
e13731554e50 undid split_comp_eq[simp] because it leads to nontermination together with split_def!
nipkow
parents: 14208
diff changeset
   550
lemma split_comp_eq: 
14101
d25c23e46173 added upd_fst, upd_snd, some thms
oheimb
parents: 13480
diff changeset
   551
"(%u. f (g (fst u)) (snd u)) = (split (%x. f (g x)))"
d25c23e46173 added upd_fst, upd_snd, some thms
oheimb
parents: 13480
diff changeset
   552
by (rule ext, auto)
d25c23e46173 added upd_fst, upd_snd, some thms
oheimb
parents: 13480
diff changeset
   553
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   554
lemma The_split_eq [simp]: "(THE (x',y'). x = x' & y = y') = (x, y)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   555
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   556
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   557
(*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   558
the following  would be slightly more general,
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   559
but cannot be used as rewrite rule:
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   560
### Cannot add premise as rewrite rule because it contains (type) unknowns:
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   561
### ?y = .x
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   562
Goal "[| P y; !!x. P x ==> x = y |] ==> (@(x',y). x = x' & P y) = (x,y)"
14208
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   563
by (rtac some_equality 1)
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   564
by ( Simp_tac 1)
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   565
by (split_all_tac 1)
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   566
by (Asm_full_simp_tac 1)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   567
qed "The_split_eq";
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   568
*)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   569
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   570
lemma injective_fst_snd: "!!x y. [|fst x = fst y; snd x = snd y|] ==> x = y"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   571
  by auto
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   572
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   573
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   574
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   575
  \bigskip @{term prod_fun} --- action of the product functor upon
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   576
  functions.
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   577
*}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   578
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   579
lemma prod_fun [simp]: "prod_fun f g (a, b) = (f a, g b)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   580
  by (simp add: prod_fun_def)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   581
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   582
lemma prod_fun_compose: "prod_fun (f1 o f2) (g1 o g2) = (prod_fun f1 g1 o prod_fun f2 g2)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   583
  apply (rule ext)
14208
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   584
  apply (tactic {* pair_tac "x" 1 *}, simp)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   585
  done
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   586
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   587
lemma prod_fun_ident [simp]: "prod_fun (%x. x) (%y. y) = (%z. z)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   588
  apply (rule ext)
14208
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   589
  apply (tactic {* pair_tac "z" 1 *}, simp)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   590
  done
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   591
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   592
lemma prod_fun_imageI [intro]: "(a, b) : r ==> (f a, g b) : prod_fun f g ` r"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   593
  apply (rule image_eqI)
14208
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   594
  apply (rule prod_fun [symmetric], assumption)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   595
  done
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   596
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   597
lemma prod_fun_imageE [elim!]:
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   598
  "[| c: (prod_fun f g)`r;  !!x y. [| c=(f(x),g(y));  (x,y):r |] ==> P
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   599
    |] ==> P"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   600
proof -
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   601
  case rule_context
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   602
  assume major: "c: (prod_fun f g)`r"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   603
  show ?thesis
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   604
    apply (rule major [THEN imageE])
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   605
    apply (rule_tac p = x in PairE)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   606
    apply (rule rule_context)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   607
     prefer 2
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   608
     apply blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   609
    apply (blast intro: prod_fun)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   610
    done
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   611
qed
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   612
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   613
14101
d25c23e46173 added upd_fst, upd_snd, some thms
oheimb
parents: 13480
diff changeset
   614
constdefs
d25c23e46173 added upd_fst, upd_snd, some thms
oheimb
parents: 13480
diff changeset
   615
  upd_fst :: "('a => 'c) => 'a * 'b => 'c * 'b"
d25c23e46173 added upd_fst, upd_snd, some thms
oheimb
parents: 13480
diff changeset
   616
 "upd_fst f == prod_fun f id"
d25c23e46173 added upd_fst, upd_snd, some thms
oheimb
parents: 13480
diff changeset
   617
d25c23e46173 added upd_fst, upd_snd, some thms
oheimb
parents: 13480
diff changeset
   618
  upd_snd :: "('b => 'c) => 'a * 'b => 'a * 'c"
d25c23e46173 added upd_fst, upd_snd, some thms
oheimb
parents: 13480
diff changeset
   619
 "upd_snd f == prod_fun id f"
d25c23e46173 added upd_fst, upd_snd, some thms
oheimb
parents: 13480
diff changeset
   620
d25c23e46173 added upd_fst, upd_snd, some thms
oheimb
parents: 13480
diff changeset
   621
lemma upd_fst_conv [simp]: "upd_fst f (x,y) = (f x,y)" 
d25c23e46173 added upd_fst, upd_snd, some thms
oheimb
parents: 13480
diff changeset
   622
by (simp add: upd_fst_def)
d25c23e46173 added upd_fst, upd_snd, some thms
oheimb
parents: 13480
diff changeset
   623
d25c23e46173 added upd_fst, upd_snd, some thms
oheimb
parents: 13480
diff changeset
   624
lemma upd_snd_conv [simp]: "upd_snd f (x,y) = (x,f y)" 
d25c23e46173 added upd_fst, upd_snd, some thms
oheimb
parents: 13480
diff changeset
   625
by (simp add: upd_snd_def)
d25c23e46173 added upd_fst, upd_snd, some thms
oheimb
parents: 13480
diff changeset
   626
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   627
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   628
  \bigskip Disjoint union of a family of sets -- Sigma.
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   629
*}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   630
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   631
lemma SigmaI [intro!]: "[| a:A;  b:B(a) |] ==> (a,b) : Sigma A B"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   632
  by (unfold Sigma_def) blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   633
14952
47455995693d removal of x-symbol syntax <Sigma> for dependent products
paulson
parents: 14565
diff changeset
   634
lemma SigmaE [elim!]:
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   635
    "[| c: Sigma A B;
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   636
        !!x y.[| x:A;  y:B(x);  c=(x,y) |] ==> P
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   637
     |] ==> P"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   638
  -- {* The general elimination rule. *}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   639
  by (unfold Sigma_def) blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   640
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   641
text {*
15422
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   642
  Elimination of @{term "(a, b) : A \<times> B"} -- introduces no
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   643
  eigenvariables.
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   644
*}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   645
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   646
lemma SigmaD1: "(a, b) : Sigma A B ==> a : A"
14952
47455995693d removal of x-symbol syntax <Sigma> for dependent products
paulson
parents: 14565
diff changeset
   647
by blast
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   648
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   649
lemma SigmaD2: "(a, b) : Sigma A B ==> b : B a"
14952
47455995693d removal of x-symbol syntax <Sigma> for dependent products
paulson
parents: 14565
diff changeset
   650
by blast
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   651
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   652
lemma SigmaE2:
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   653
    "[| (a, b) : Sigma A B;
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   654
        [| a:A;  b:B(a) |] ==> P
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   655
     |] ==> P"
14952
47455995693d removal of x-symbol syntax <Sigma> for dependent products
paulson
parents: 14565
diff changeset
   656
  by blast
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   657
14952
47455995693d removal of x-symbol syntax <Sigma> for dependent products
paulson
parents: 14565
diff changeset
   658
lemma Sigma_cong:
15422
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   659
     "\<lbrakk>A = B; !!x. x \<in> B \<Longrightarrow> C x = D x\<rbrakk>
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   660
      \<Longrightarrow> (SIGMA x: A. C x) = (SIGMA x: B. D x)"
14952
47455995693d removal of x-symbol syntax <Sigma> for dependent products
paulson
parents: 14565
diff changeset
   661
by auto
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   662
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   663
lemma Sigma_mono: "[| A <= C; !!x. x:A ==> B x <= D x |] ==> Sigma A B <= Sigma C D"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   664
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   665
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   666
lemma Sigma_empty1 [simp]: "Sigma {} B = {}"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   667
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   668
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   669
lemma Sigma_empty2 [simp]: "A <*> {} = {}"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   670
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   671
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   672
lemma UNIV_Times_UNIV [simp]: "UNIV <*> UNIV = UNIV"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   673
  by auto
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   674
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   675
lemma Compl_Times_UNIV1 [simp]: "- (UNIV <*> A) = UNIV <*> (-A)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   676
  by auto
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   677
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   678
lemma Compl_Times_UNIV2 [simp]: "- (A <*> UNIV) = (-A) <*> UNIV"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   679
  by auto
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   680
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   681
lemma mem_Sigma_iff [iff]: "((a,b): Sigma A B) = (a:A & b:B(a))"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   682
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   683
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   684
lemma Times_subset_cancel2: "x:C ==> (A <*> C <= B <*> C) = (A <= B)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   685
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   686
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   687
lemma Times_eq_cancel2: "x:C ==> (A <*> C = B <*> C) = (A = B)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   688
  by (blast elim: equalityE)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   689
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   690
lemma SetCompr_Sigma_eq:
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   691
    "Collect (split (%x y. P x & Q x y)) = (SIGMA x:Collect P. Collect (Q x))"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   692
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   693
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   694
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   695
  \bigskip Complex rules for Sigma.
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   696
*}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   697
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   698
lemma Collect_split [simp]: "{(a,b). P a & Q b} = Collect P <*> Collect Q"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   699
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   700
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   701
lemma UN_Times_distrib:
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   702
  "(UN (a,b):(A <*> B). E a <*> F b) = (UNION A E) <*> (UNION B F)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   703
  -- {* Suggested by Pierre Chartier *}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   704
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   705
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   706
lemma split_paired_Ball_Sigma [simp]:
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   707
    "(ALL z: Sigma A B. P z) = (ALL x:A. ALL y: B x. P(x,y))"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   708
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   709
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   710
lemma split_paired_Bex_Sigma [simp]:
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   711
    "(EX z: Sigma A B. P z) = (EX x:A. EX y: B x. P(x,y))"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   712
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   713
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   714
lemma Sigma_Un_distrib1: "(SIGMA i:I Un J. C(i)) = (SIGMA i:I. C(i)) Un (SIGMA j:J. C(j))"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   715
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   716
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   717
lemma Sigma_Un_distrib2: "(SIGMA i:I. A(i) Un B(i)) = (SIGMA i:I. A(i)) Un (SIGMA i:I. B(i))"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   718
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   719
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   720
lemma Sigma_Int_distrib1: "(SIGMA i:I Int J. C(i)) = (SIGMA i:I. C(i)) Int (SIGMA j:J. C(j))"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   721
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   722
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   723
lemma Sigma_Int_distrib2: "(SIGMA i:I. A(i) Int B(i)) = (SIGMA i:I. A(i)) Int (SIGMA i:I. B(i))"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   724
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   725
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   726
lemma Sigma_Diff_distrib1: "(SIGMA i:I - J. C(i)) = (SIGMA i:I. C(i)) - (SIGMA j:J. C(j))"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   727
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   728
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   729
lemma Sigma_Diff_distrib2: "(SIGMA i:I. A(i) - B(i)) = (SIGMA i:I. A(i)) - (SIGMA i:I. B(i))"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   730
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   731
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   732
lemma Sigma_Union: "Sigma (Union X) B = (UN A:X. Sigma A B)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   733
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   734
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   735
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   736
  Non-dependent versions are needed to avoid the need for higher-order
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   737
  matching, especially when the rules are re-oriented.
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   738
*}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   739
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   740
lemma Times_Un_distrib1: "(A Un B) <*> C = (A <*> C) Un (B <*> C)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   741
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   742
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   743
lemma Times_Int_distrib1: "(A Int B) <*> C = (A <*> C) Int (B <*> C)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   744
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   745
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   746
lemma Times_Diff_distrib1: "(A - B) <*> C = (A <*> C) - (B <*> C)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   747
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   748
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   749
11493
f3ff2549cdc8 added pair_imageI (also as intro rule)
oheimb
parents: 11451
diff changeset
   750
lemma pair_imageI [intro]: "(a, b) : A ==> f a b : (%(a, b). f a b) ` A"
11777
wenzelm
parents: 11602
diff changeset
   751
  apply (rule_tac x = "(a, b)" in image_eqI)
wenzelm
parents: 11602
diff changeset
   752
   apply auto
wenzelm
parents: 11602
diff changeset
   753
  done
wenzelm
parents: 11602
diff changeset
   754
11493
f3ff2549cdc8 added pair_imageI (also as intro rule)
oheimb
parents: 11451
diff changeset
   755
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   756
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   757
  Setup of internal @{text split_rule}.
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   758
*}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   759
11032
83f723e86dac added hidden internal_split constant;
wenzelm
parents: 11025
diff changeset
   760
constdefs
11425
wenzelm
parents: 11032
diff changeset
   761
  internal_split :: "('a => 'b => 'c) => 'a * 'b => 'c"
11032
83f723e86dac added hidden internal_split constant;
wenzelm
parents: 11025
diff changeset
   762
  "internal_split == split"
83f723e86dac added hidden internal_split constant;
wenzelm
parents: 11025
diff changeset
   763
83f723e86dac added hidden internal_split constant;
wenzelm
parents: 11025
diff changeset
   764
lemma internal_split_conv: "internal_split c (a, b) = c a b"
83f723e86dac added hidden internal_split constant;
wenzelm
parents: 11025
diff changeset
   765
  by (simp only: internal_split_def split_conv)
83f723e86dac added hidden internal_split constant;
wenzelm
parents: 11025
diff changeset
   766
83f723e86dac added hidden internal_split constant;
wenzelm
parents: 11025
diff changeset
   767
hide const internal_split
83f723e86dac added hidden internal_split constant;
wenzelm
parents: 11025
diff changeset
   768
11025
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   769
use "Tools/split_rule.ML"
11032
83f723e86dac added hidden internal_split constant;
wenzelm
parents: 11025
diff changeset
   770
setup SplitRule.setup
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   771
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   772
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   773
subsection {* Code generator setup *}
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   774
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   775
types_code
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   776
  "*"     ("(_ */ _)")
16770
1f1b1fae30e4 Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents: 16634
diff changeset
   777
attach (term_of) {*
1f1b1fae30e4 Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents: 16634
diff changeset
   778
fun term_of_id_42 f T g U (x, y) = HOLogic.pair_const T U $ f x $ g y;
1f1b1fae30e4 Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents: 16634
diff changeset
   779
*}
1f1b1fae30e4 Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents: 16634
diff changeset
   780
attach (test) {*
1f1b1fae30e4 Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents: 16634
diff changeset
   781
fun gen_id_42 aG bG i = (aG i, bG i);
1f1b1fae30e4 Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents: 16634
diff changeset
   782
*}
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   783
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   784
consts_code
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   785
  "Pair"    ("(_,/ _)")
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   786
  "fst"     ("fst")
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   787
  "snd"     ("snd")
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   788
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   789
ML {*
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   790
local
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   791
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   792
fun strip_abs 0 t = ([], t)
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   793
  | strip_abs i (Abs (s, T, t)) =
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   794
    let
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   795
      val s' = Codegen.new_name t s;
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   796
      val v = Free (s', T)
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   797
    in apfst (cons v) (strip_abs (i-1) (subst_bound (v, t))) end
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   798
  | strip_abs i (u as Const ("split", _) $ t) = (case strip_abs (i+1) t of
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   799
        (v :: v' :: vs, u) => (HOLogic.mk_prod (v, v') :: vs, u)
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   800
      | _ => ([], u))
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   801
  | strip_abs i t = ([], t);
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   802
16634
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   803
fun let_codegen thy defs gr dep thyname brack t = (case strip_comb t of
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   804
    (t1 as Const ("Let", _), t2 :: t3 :: ts) =>
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   805
    let
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   806
      fun dest_let (l as Const ("Let", _) $ t $ u) =
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   807
          (case strip_abs 1 u of
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   808
             ([p], u') => apfst (cons (p, t)) (dest_let u')
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   809
           | _ => ([], l))
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   810
        | dest_let t = ([], t);
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   811
      fun mk_code (gr, (l, r)) =
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   812
        let
16634
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   813
          val (gr1, pl) = Codegen.invoke_codegen thy defs dep thyname false (gr, l);
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   814
          val (gr2, pr) = Codegen.invoke_codegen thy defs dep thyname false (gr1, r);
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   815
        in (gr2, (pl, pr)) end
16634
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   816
    in case dest_let (t1 $ t2 $ t3) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   817
        ([], _) => NONE
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   818
      | (ps, u) =>
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   819
          let
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   820
            val (gr1, qs) = foldl_map mk_code (gr, ps);
16634
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   821
            val (gr2, pu) = Codegen.invoke_codegen thy defs dep thyname false (gr1, u);
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   822
            val (gr3, pargs) = foldl_map
17021
1c361a3de73d Fixed bug in code generator for let and split leading to ill-formed code.
berghofe
parents: 17002
diff changeset
   823
              (Codegen.invoke_codegen thy defs dep thyname true) (gr2, ts)
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   824
          in
16634
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   825
            SOME (gr3, Codegen.mk_app brack
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   826
              (Pretty.blk (0, [Pretty.str "let ", Pretty.blk (0, List.concat
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   827
                  (separate [Pretty.str ";", Pretty.brk 1] (map (fn (pl, pr) =>
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   828
                    [Pretty.block [Pretty.str "val ", pl, Pretty.str " =",
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   829
                       Pretty.brk 1, pr]]) qs))),
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   830
                Pretty.brk 1, Pretty.str "in ", pu,
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   831
                Pretty.brk 1, Pretty.str "end"])) pargs)
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   832
          end
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   833
    end
16634
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   834
  | _ => NONE);
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   835
16634
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   836
fun split_codegen thy defs gr dep thyname brack t = (case strip_comb t of
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   837
    (t1 as Const ("split", _), t2 :: ts) =>
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   838
      (case strip_abs 1 (t1 $ t2) of
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   839
         ([p], u) =>
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   840
           let
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   841
             val (gr1, q) = Codegen.invoke_codegen thy defs dep thyname false (gr, p);
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   842
             val (gr2, pu) = Codegen.invoke_codegen thy defs dep thyname false (gr1, u);
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   843
             val (gr3, pargs) = foldl_map
17021
1c361a3de73d Fixed bug in code generator for let and split leading to ill-formed code.
berghofe
parents: 17002
diff changeset
   844
               (Codegen.invoke_codegen thy defs dep thyname true) (gr2, ts)
16634
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   845
           in
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   846
             SOME (gr2, Codegen.mk_app brack
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   847
               (Pretty.block [Pretty.str "(fn ", q, Pretty.str " =>",
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   848
                 Pretty.brk 1, pu, Pretty.str ")"]) pargs)
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   849
           end
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   850
       | _ => NONE)
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   851
  | _ => NONE);
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   852
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   853
in
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   854
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   855
val prod_codegen_setup =
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   856
  [Codegen.add_codegen "let_codegen" let_codegen,
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   857
   Codegen.add_codegen "split_codegen" split_codegen];
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   858
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   859
end;
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   860
*}
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   861
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   862
setup prod_codegen_setup
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   863
15404
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   864
ML
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   865
{*
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   866
val Collect_split = thm "Collect_split";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   867
val Compl_Times_UNIV1 = thm "Compl_Times_UNIV1";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   868
val Compl_Times_UNIV2 = thm "Compl_Times_UNIV2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   869
val PairE = thm "PairE";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   870
val PairE_lemma = thm "PairE_lemma";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   871
val Pair_Rep_inject = thm "Pair_Rep_inject";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   872
val Pair_def = thm "Pair_def";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   873
val Pair_eq = thm "Pair_eq";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   874
val Pair_fst_snd_eq = thm "Pair_fst_snd_eq";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   875
val Pair_inject = thm "Pair_inject";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   876
val ProdI = thm "ProdI";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   877
val SetCompr_Sigma_eq = thm "SetCompr_Sigma_eq";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   878
val SigmaD1 = thm "SigmaD1";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   879
val SigmaD2 = thm "SigmaD2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   880
val SigmaE = thm "SigmaE";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   881
val SigmaE2 = thm "SigmaE2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   882
val SigmaI = thm "SigmaI";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   883
val Sigma_Diff_distrib1 = thm "Sigma_Diff_distrib1";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   884
val Sigma_Diff_distrib2 = thm "Sigma_Diff_distrib2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   885
val Sigma_Int_distrib1 = thm "Sigma_Int_distrib1";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   886
val Sigma_Int_distrib2 = thm "Sigma_Int_distrib2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   887
val Sigma_Un_distrib1 = thm "Sigma_Un_distrib1";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   888
val Sigma_Un_distrib2 = thm "Sigma_Un_distrib2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   889
val Sigma_Union = thm "Sigma_Union";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   890
val Sigma_def = thm "Sigma_def";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   891
val Sigma_empty1 = thm "Sigma_empty1";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   892
val Sigma_empty2 = thm "Sigma_empty2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   893
val Sigma_mono = thm "Sigma_mono";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   894
val The_split = thm "The_split";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   895
val The_split_eq = thm "The_split_eq";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   896
val The_split_eq = thm "The_split_eq";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   897
val Times_Diff_distrib1 = thm "Times_Diff_distrib1";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   898
val Times_Int_distrib1 = thm "Times_Int_distrib1";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   899
val Times_Un_distrib1 = thm "Times_Un_distrib1";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   900
val Times_eq_cancel2 = thm "Times_eq_cancel2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   901
val Times_subset_cancel2 = thm "Times_subset_cancel2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   902
val UNIV_Times_UNIV = thm "UNIV_Times_UNIV";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   903
val UN_Times_distrib = thm "UN_Times_distrib";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   904
val Unity_def = thm "Unity_def";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   905
val cond_split_eta = thm "cond_split_eta";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   906
val fst_conv = thm "fst_conv";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   907
val fst_def = thm "fst_def";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   908
val fst_eqD = thm "fst_eqD";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   909
val inj_on_Abs_Prod = thm "inj_on_Abs_Prod";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   910
val injective_fst_snd = thm "injective_fst_snd";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   911
val mem_Sigma_iff = thm "mem_Sigma_iff";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   912
val mem_splitE = thm "mem_splitE";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   913
val mem_splitI = thm "mem_splitI";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   914
val mem_splitI2 = thm "mem_splitI2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   915
val prod_eqI = thm "prod_eqI";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   916
val prod_fun = thm "prod_fun";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   917
val prod_fun_compose = thm "prod_fun_compose";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   918
val prod_fun_def = thm "prod_fun_def";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   919
val prod_fun_ident = thm "prod_fun_ident";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   920
val prod_fun_imageE = thm "prod_fun_imageE";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   921
val prod_fun_imageI = thm "prod_fun_imageI";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   922
val prod_induct = thm "prod_induct";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   923
val snd_conv = thm "snd_conv";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   924
val snd_def = thm "snd_def";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   925
val snd_eqD = thm "snd_eqD";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   926
val split = thm "split";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   927
val splitD = thm "splitD";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   928
val splitD' = thm "splitD'";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   929
val splitE = thm "splitE";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   930
val splitE' = thm "splitE'";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   931
val splitE2 = thm "splitE2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   932
val splitI = thm "splitI";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   933
val splitI2 = thm "splitI2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   934
val splitI2' = thm "splitI2'";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   935
val split_Pair_apply = thm "split_Pair_apply";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   936
val split_beta = thm "split_beta";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   937
val split_conv = thm "split_conv";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   938
val split_def = thm "split_def";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   939
val split_eta = thm "split_eta";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   940
val split_eta_SetCompr = thm "split_eta_SetCompr";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   941
val split_eta_SetCompr2 = thm "split_eta_SetCompr2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   942
val split_paired_All = thm "split_paired_All";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   943
val split_paired_Ball_Sigma = thm "split_paired_Ball_Sigma";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   944
val split_paired_Bex_Sigma = thm "split_paired_Bex_Sigma";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   945
val split_paired_Ex = thm "split_paired_Ex";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   946
val split_paired_The = thm "split_paired_The";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   947
val split_paired_all = thm "split_paired_all";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   948
val split_part = thm "split_part";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   949
val split_split = thm "split_split";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   950
val split_split_asm = thm "split_split_asm";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   951
val split_tupled_all = thms "split_tupled_all";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   952
val split_weak_cong = thm "split_weak_cong";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   953
val surj_pair = thm "surj_pair";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   954
val surjective_pairing = thm "surjective_pairing";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   955
val unit_abs_eta_conv = thm "unit_abs_eta_conv";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   956
val unit_all_eq1 = thm "unit_all_eq1";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   957
val unit_all_eq2 = thm "unit_all_eq2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   958
val unit_eq = thm "unit_eq";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   959
val unit_induct = thm "unit_induct";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   960
*}
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
   961
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   962
end