src/HOL/Product_Type.thy
author berghofe
Mon, 26 Oct 2009 14:53:33 +0100
changeset 33189 82a40677c1f8
parent 33089 4e33c819fced
child 33271 7be66dee1a5a
permissions -rw-r--r--
Added Pattern.thy to Nominal/Examples.
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
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     3
    Copyright   1992  University of Cambridge
11777
wenzelm
parents: 11602
diff changeset
     4
*)
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     5
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
     6
header {* Cartesian products *}
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     7
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14952
diff changeset
     8
theory Product_Type
33089
4e33c819fced import theory Nat here, which avoids duplicate definition of datatype_realizers (and thus allows to maintain fully authentic fact table);
wenzelm
parents: 32952
diff changeset
     9
imports Inductive Nat
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    10
uses
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    11
  ("Tools/split_rule.ML")
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31667
diff changeset
    12
  ("Tools/inductive_set.ML")
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    13
  ("Tools/inductive_realizer.ML")
31775
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 31723
diff changeset
    14
  ("Tools/Datatype/datatype_realizer.ML")
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14952
diff changeset
    15
begin
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    16
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    17
subsection {* @{typ bool} is a datatype *}
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    18
27104
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 26975
diff changeset
    19
rep_datatype True False by (auto intro: bool_induct)
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    20
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    21
declare case_split [cases type: bool]
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    22
  -- "prefer plain propositional version"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    23
28346
b8390cd56b8f discontinued special treatment of op = vs. eq_class.eq
haftmann
parents: 28262
diff changeset
    24
lemma
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28537
diff changeset
    25
  shows [code]: "eq_class.eq False P \<longleftrightarrow> \<not> P"
4e74209f113e `code func` now just `code`
haftmann
parents: 28537
diff changeset
    26
    and [code]: "eq_class.eq True P \<longleftrightarrow> P" 
4e74209f113e `code func` now just `code`
haftmann
parents: 28537
diff changeset
    27
    and [code]: "eq_class.eq P False \<longleftrightarrow> \<not> P" 
4e74209f113e `code func` now just `code`
haftmann
parents: 28537
diff changeset
    28
    and [code]: "eq_class.eq P True \<longleftrightarrow> P"
28346
b8390cd56b8f discontinued special treatment of op = vs. eq_class.eq
haftmann
parents: 28262
diff changeset
    29
    and [code nbe]: "eq_class.eq P P \<longleftrightarrow> True"
b8390cd56b8f discontinued special treatment of op = vs. eq_class.eq
haftmann
parents: 28262
diff changeset
    30
  by (simp_all add: eq)
25534
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25511
diff changeset
    31
28346
b8390cd56b8f discontinued special treatment of op = vs. eq_class.eq
haftmann
parents: 28262
diff changeset
    32
code_const "eq_class.eq \<Colon> bool \<Rightarrow> bool \<Rightarrow> bool"
25534
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25511
diff changeset
    33
  (Haskell infixl 4 "==")
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25511
diff changeset
    34
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25511
diff changeset
    35
code_instance bool :: eq
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25511
diff changeset
    36
  (Haskell -)
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    37
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
    38
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    39
subsection {* Unit *}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    40
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    41
typedef unit = "{True}"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    42
proof
20588
c847c56edf0c added operational equality
haftmann
parents: 20415
diff changeset
    43
  show "True : ?unit" ..
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    44
qed
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    45
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    46
definition
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    47
  Unity :: unit    ("'(')")
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    48
where
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    49
  "() = Abs_unit True"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    50
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24162
diff changeset
    51
lemma unit_eq [noatp]: "u = ()"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    52
  by (induct u) (simp add: unit_def Unity_def)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    53
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    54
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    55
  Simplification procedure for @{thm [source] unit_eq}.  Cannot use
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    56
  this rule directly --- it loops!
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
26480
544cef16045b replaced 'ML_setup' by 'ML';
wenzelm
parents: 26358
diff changeset
    59
ML {*
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12338
diff changeset
    60
  val unit_eq_proc =
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    61
    let val unit_meta_eq = mk_meta_eq @{thm unit_eq} in
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    62
      Simplifier.simproc @{theory} "unit_eq" ["x::unit"]
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
    63
      (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
    64
    end;
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    65
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    66
  Addsimprocs [unit_eq_proc];
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    67
*}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    68
27104
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 26975
diff changeset
    69
rep_datatype "()" by simp
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    70
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    71
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
    72
  by simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    73
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    74
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
    75
  by (rule triv_forall_equality)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    76
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    77
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    78
  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
    79
  [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
    80
  f} rather than by @{term [source] "%u. f ()"}.
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    81
*}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    82
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24162
diff changeset
    83
lemma unit_abs_eta_conv [simp,noatp]: "(%u::unit. f ()) = f"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    84
  by (rule ext) simp
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    85
30924
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
    86
instantiation unit :: default
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
    87
begin
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
    88
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
    89
definition "default = ()"
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
    90
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
    91
instance ..
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
    92
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
    93
end
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    94
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
    95
text {* code generator setup *}
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
    96
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
    97
instance unit :: eq ..
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
    98
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28537
diff changeset
    99
lemma [code]:
28346
b8390cd56b8f discontinued special treatment of op = vs. eq_class.eq
haftmann
parents: 28262
diff changeset
   100
  "eq_class.eq (u\<Colon>unit) v \<longleftrightarrow> True" unfolding eq unit_eq [of u] unit_eq [of v] by rule+
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   101
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   102
code_type unit
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   103
  (SML "unit")
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   104
  (OCaml "unit")
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   105
  (Haskell "()")
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   106
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   107
code_instance unit :: eq
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   108
  (Haskell -)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   109
28346
b8390cd56b8f discontinued special treatment of op = vs. eq_class.eq
haftmann
parents: 28262
diff changeset
   110
code_const "eq_class.eq \<Colon> unit \<Rightarrow> unit \<Rightarrow> bool"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   111
  (Haskell infixl 4 "==")
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   112
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   113
code_const Unity
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   114
  (SML "()")
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   115
  (OCaml "()")
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   116
  (Haskell "()")
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   117
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   118
code_reserved SML
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   119
  unit
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   120
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   121
code_reserved OCaml
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   122
  unit
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   123
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   124
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   125
subsection {* Pairs *}
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   126
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   127
subsubsection {* Product type, basic operations and concrete syntax *}
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   128
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   129
definition
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   130
  Pair_Rep :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   131
where
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   132
  "Pair_Rep a b = (\<lambda>x y. x = a \<and> y = b)"
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   133
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   134
global
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   135
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   136
typedef (Prod)
22838
haftmann
parents: 22744
diff changeset
   137
  ('a, 'b) "*"    (infixr "*" 20)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   138
    = "{f. \<exists>a b. f = Pair_Rep (a\<Colon>'a) (b\<Colon>'b)}"
11025
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   139
proof
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   140
  fix a b show "Pair_Rep a b \<in> ?Prod"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   141
    by rule+
11025
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   142
qed
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   143
12114
a8e860c86252 eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents: 11966
diff changeset
   144
syntax (xsymbols)
15422
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   145
  "*"      :: "[type, type] => type"         ("(_ \<times>/ _)" [21, 20] 20)
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   146
syntax (HTML output)
15422
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   147
  "*"      :: "[type, type] => type"         ("(_ \<times>/ _)" [21, 20] 20)
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   148
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   149
consts
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   150
  Pair     :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<times> 'b"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   151
  fst      :: "'a \<times> 'b \<Rightarrow> 'a"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   152
  snd      :: "'a \<times> 'b \<Rightarrow> 'b"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   153
  split    :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   154
  curry    :: "('a \<times> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'c"
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   155
11777
wenzelm
parents: 11602
diff changeset
   156
local
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   157
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19179
diff changeset
   158
defs
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19179
diff changeset
   159
  Pair_def:     "Pair a b == Abs_Prod (Pair_Rep a b)"
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19179
diff changeset
   160
  fst_def:      "fst p == THE a. EX b. p = Pair a b"
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19179
diff changeset
   161
  snd_def:      "snd p == THE b. EX a. p = Pair a b"
24162
8dfd5dd65d82 split off theory Option for benefit of code generator
haftmann
parents: 23247
diff changeset
   162
  split_def:    "split == (%c p. c (fst p) (snd p))"
8dfd5dd65d82 split off theory Option for benefit of code generator
haftmann
parents: 23247
diff changeset
   163
  curry_def:    "curry == (%c x y. c (Pair x y))"
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19179
diff changeset
   164
11777
wenzelm
parents: 11602
diff changeset
   165
text {*
wenzelm
parents: 11602
diff changeset
   166
  Patterns -- extends pre-defined type @{typ pttrn} used in
wenzelm
parents: 11602
diff changeset
   167
  abstractions.
wenzelm
parents: 11602
diff changeset
   168
*}
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   169
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   170
nonterminals
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   171
  tuple_args patterns
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   172
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   173
syntax
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   174
  "_tuple"      :: "'a => tuple_args => 'a * 'b"        ("(1'(_,/ _'))")
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   175
  "_tuple_arg"  :: "'a => tuple_args"                   ("_")
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   176
  "_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
   177
  "_pattern"    :: "[pttrn, patterns] => pttrn"         ("'(_,/ _')")
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   178
  ""            :: "pttrn => patterns"                  ("_")
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   179
  "_patterns"   :: "[pttrn, patterns] => patterns"      ("_,/ _")
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   180
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   181
translations
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   182
  "(x, y)"       == "Pair x y"
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   183
  "_tuple x (_tuple_args y z)" == "_tuple x (_tuple_arg (_tuple y z))"
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   184
  "%(x,y,zs).b"  == "split(%x (y,zs).b)"
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   185
  "%(x,y).b"     == "split(%x y. b)"
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   186
  "_abs (Pair x y) t" => "%(x,y).t"
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   187
  (* The last rule accommodates tuples in `case C ... (x,y) ... => ...'
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   188
     The (x,y) is parsed as `Pair x y' because it is logic, not pttrn *)
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   189
14359
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   190
(* reconstructs pattern from (nested) splits, avoiding eta-contraction of body*)
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   191
(* works best with enclosing "let", if "let" does not avoid eta-contraction   *)
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   192
print_translation {*
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   193
let fun split_tr' [Abs (x,T,t as (Abs abs))] =
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   194
      (* split (%x y. t) => %(x,y) t *)
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   195
      let val (y,t') = atomic_abs_tr' abs;
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   196
          val (x',t'') = atomic_abs_tr' (x,T,t');
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   197
    
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   198
      in Syntax.const "_abs" $ (Syntax.const "_pattern" $x'$y) $ t'' end
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   199
    | split_tr' [Abs (x,T,(s as Const ("split",_)$t))] =
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   200
       (* split (%x. (split (%y z. t))) => %(x,y,z). t *)
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   201
       let val (Const ("_abs",_)$(Const ("_pattern",_)$y$z)$t') = split_tr' [t];
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   202
           val (x',t'') = atomic_abs_tr' (x,T,t');
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   203
       in Syntax.const "_abs"$ 
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   204
           (Syntax.const "_pattern"$x'$(Syntax.const "_patterns"$y$z))$t'' end
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   205
    | split_tr' [Const ("split",_)$t] =
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   206
       (* split (split (%x y z. t)) => %((x,y),z). t *)   
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   207
       split_tr' [(split_tr' [t])] (* inner split_tr' creates next pattern *)
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   208
    | split_tr' [Const ("_abs",_)$x_y$(Abs abs)] =
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   209
       (* split (%pttrn z. t) => %(pttrn,z). t *)
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   210
       let val (z,t) = atomic_abs_tr' abs;
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   211
       in Syntax.const "_abs" $ (Syntax.const "_pattern" $x_y$z) $ t end
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   212
    | split_tr' _ =  raise Match;
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   213
in [("split", split_tr')]
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   214
end
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   215
*}
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   216
15422
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   217
(* 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
   218
typed_print_translation {*
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   219
let
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   220
  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
   221
    | 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
   222
        (case (head_of t) of
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   223
           Const ("split",_) => raise Match
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   224
         | _ => let 
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   225
                  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
   226
                  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
   227
                  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
   228
                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
   229
    | split_guess_names_tr' _ T [t] =
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   230
       (case (head_of t) of
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   231
           Const ("split",_) => raise Match 
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   232
         | _ => let 
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   233
                  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
   234
                  val (y,t') = 
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   235
                        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
   236
                  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
   237
                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
   238
    | split_guess_names_tr' _ _ _ = raise Match;
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   239
in [("split", split_guess_names_tr')]
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   240
end 
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   241
*}
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   242
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   243
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   244
text {* Towards a datatype declaration *}
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   245
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   246
lemma surj_pair [simp]: "EX x y. p = (x, y)"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   247
  apply (unfold Pair_def)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   248
  apply (rule Rep_Prod [unfolded Prod_def, THEN CollectE])
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   249
  apply (erule exE, erule exE, rule exI, rule exI)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   250
  apply (rule Rep_Prod_inverse [symmetric, THEN trans])
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   251
  apply (erule arg_cong)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   252
  done
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   253
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   254
lemma PairE [cases type: *]:
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   255
  obtains x y where "p = (x, y)"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   256
  using surj_pair [of p] by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   257
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   258
lemma ProdI: "Pair_Rep a b \<in> Prod"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   259
  unfolding Prod_def by rule+
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   260
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   261
lemma Pair_Rep_inject: "Pair_Rep a b = Pair_Rep a' b' \<Longrightarrow> a = a' \<and> b = b'"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   262
  unfolding Pair_Rep_def by (drule fun_cong, drule fun_cong) blast
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   263
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   264
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
   265
  apply (rule inj_on_inverseI)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   266
  apply (erule Abs_Prod_inverse)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   267
  done
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   268
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   269
lemma Pair_inject:
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   270
  assumes "(a, b) = (a', b')"
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   271
    and "a = a' ==> b = b' ==> R"
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   272
  shows R
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   273
  apply (insert prems [unfolded Pair_def])
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   274
  apply (rule inj_on_Abs_Prod [THEN inj_onD, THEN Pair_Rep_inject, THEN conjE])
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   275
  apply (assumption | rule ProdI)+
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   276
  done
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   277
27104
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 26975
diff changeset
   278
rep_datatype (prod) Pair
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 26975
diff changeset
   279
proof -
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 26975
diff changeset
   280
  fix P p
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 26975
diff changeset
   281
  assume "\<And>x y. P (x, y)"
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 26975
diff changeset
   282
  then show "P p" by (cases p) simp
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 26975
diff changeset
   283
qed (auto elim: Pair_inject)
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 26975
diff changeset
   284
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 26975
diff changeset
   285
lemmas Pair_eq = prod.inject
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   286
22886
cdff6ef76009 moved recfun_codegen.ML to Code_Generator.thy
haftmann
parents: 22838
diff changeset
   287
lemma fst_conv [simp, code]: "fst (a, b) = a"
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19179
diff changeset
   288
  unfolding fst_def by blast
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   289
22886
cdff6ef76009 moved recfun_codegen.ML to Code_Generator.thy
haftmann
parents: 22838
diff changeset
   290
lemma snd_conv [simp, code]: "snd (a, b) = b"
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19179
diff changeset
   291
  unfolding snd_def by blast
11025
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   292
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   293
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   294
subsubsection {* Basic rules and proof tools *}
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   295
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   296
lemma fst_eqD: "fst (x, y) = a ==> x = a"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   297
  by simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   298
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   299
lemma snd_eqD: "snd (x, y) = a ==> y = a"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   300
  by simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   301
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   302
lemma pair_collapse [simp]: "(fst p, snd p) = p"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   303
  by (cases p) simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   304
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   305
lemmas surjective_pairing = pair_collapse [symmetric]
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   306
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   307
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
   308
proof
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   309
  fix a b
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   310
  assume "!!x. PROP P x"
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19179
diff changeset
   311
  then show "PROP P (a, b)" .
11820
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   312
next
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   313
  fix x
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   314
  assume "!!a b. PROP P (a, b)"
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19179
diff changeset
   315
  from `PROP P (fst x, snd x)` show "PROP P x" by simp
11820
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   316
qed
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   317
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   318
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   319
  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
   320
  Simplifier because it also affects premises in congrence rules,
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   321
  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
   322
  ?P(a, b)"} which cannot be solved by reflexivity.
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   323
*}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   324
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   325
lemmas split_tupled_all = split_paired_all unit_all_eq2
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   326
26480
544cef16045b replaced 'ML_setup' by 'ML';
wenzelm
parents: 26358
diff changeset
   327
ML {*
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   328
  (* replace parameters of product type by individual component parameters *)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   329
  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
   330
  local (* filtering with exists_paired_all is an essential optimization *)
16121
wenzelm
parents: 15570
diff changeset
   331
    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
   332
          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
   333
      | 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
   334
      | exists_paired_all (Abs (_, _, t)) = exists_paired_all t
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   335
      | exists_paired_all _ = false;
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   336
    val ss = HOL_basic_ss
26340
a85fe32e7b2f more antiquotations;
wenzelm
parents: 26143
diff changeset
   337
      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
   338
      addsimprocs [unit_eq_proc];
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   339
  in
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   340
    val split_all_tac = SUBGOAL (fn (t, i) =>
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   341
      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
   342
    val unsafe_split_all_tac = SUBGOAL (fn (t, i) =>
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   343
      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
   344
    fun split_all th =
26340
a85fe32e7b2f more antiquotations;
wenzelm
parents: 26143
diff changeset
   345
   if exists_paired_all (Thm.prop_of th) then full_simplify ss th else th;
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   346
  end;
26340
a85fe32e7b2f more antiquotations;
wenzelm
parents: 26143
diff changeset
   347
*}
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   348
26340
a85fe32e7b2f more antiquotations;
wenzelm
parents: 26143
diff changeset
   349
declaration {* fn _ =>
a85fe32e7b2f more antiquotations;
wenzelm
parents: 26143
diff changeset
   350
  Classical.map_cs (fn cs => cs addSbefore ("split_all_tac", split_all_tac))
16121
wenzelm
parents: 15570
diff changeset
   351
*}
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   352
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   353
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
   354
  -- {* @{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
   355
  by fast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   356
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   357
lemma split_paired_Ex [simp]: "(EX x. P x) = (EX a b. P (a, b))"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   358
  by fast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   359
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   360
lemma Pair_fst_snd_eq: "s = t \<longleftrightarrow> fst s = fst t \<and> snd s = snd t"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   361
  by (cases s, cases t) simp
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   362
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   363
lemma prod_eqI [intro?]: "fst p = fst q \<Longrightarrow> snd p = snd q \<Longrightarrow> p = q"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   364
  by (simp add: Pair_fst_snd_eq)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   365
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   366
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   367
subsubsection {* @{text split} and @{text curry} *}
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   368
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28537
diff changeset
   369
lemma split_conv [simp, code]: "split f (a, b) = f a b"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   370
  by (simp add: split_def)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   371
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28537
diff changeset
   372
lemma curry_conv [simp, code]: "curry f a b = f (a, b)"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   373
  by (simp add: curry_def)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   374
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   375
lemmas split = split_conv  -- {* for backwards compatibility *}
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   376
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   377
lemma splitI: "f a b \<Longrightarrow> split f (a, b)"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   378
  by (rule split_conv [THEN iffD2])
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   379
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   380
lemma splitD: "split f (a, b) \<Longrightarrow> f a b"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   381
  by (rule split_conv [THEN iffD1])
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   382
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   383
lemma curryI [intro!]: "f (a, b) \<Longrightarrow> curry f a b"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   384
  by (simp add: curry_def)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   385
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   386
lemma curryD [dest!]: "curry f a b \<Longrightarrow> f (a, b)"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   387
  by (simp add: curry_def)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   388
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   389
lemma curryE: "curry f a b \<Longrightarrow> (f (a, b) \<Longrightarrow> Q) \<Longrightarrow> Q"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   390
  by (simp add: curry_def)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   391
14189
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
   392
lemma curry_split [simp]: "curry (split f) = f"
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
   393
  by (simp add: curry_def split_def)
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
   394
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
   395
lemma split_curry [simp]: "split (curry f) = f"
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
   396
  by (simp add: curry_def split_def)
de58f4d939e1 Added the constant "curry".
skalberg
parents: 14101
diff changeset
   397
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   398
lemma split_Pair [simp]: "(\<lambda>(x, y). (x, y)) = id"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   399
  by (simp add: split_def id_def)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   400
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   401
lemma split_eta: "(\<lambda>(x, y). f (x, y)) = f"
31775
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 31723
diff changeset
   402
  -- {* Subsumes the old @{text split_Pair} when @{term f} is the identity Datatype. *}
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   403
  by (rule ext) auto
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   404
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   405
lemma split_comp: "split (f \<circ> g) x = f (g (fst x)) (snd x)"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   406
  by (cases x) simp
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   407
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   408
lemma split_twice: "split f (split g p) = split (\<lambda>x y. split f (g x y)) p"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   409
  unfolding split_def ..
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   410
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   411
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
   412
  -- {* Can't be added to simpset: loops! *}
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   413
  by (simp add: split_eta)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   414
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   415
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
   416
  by (simp add: split_def)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   417
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   418
lemma split_weak_cong: "p = q \<Longrightarrow> split c p = split c q"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   419
  -- {* Prevents simplification of @{term c}: much faster *}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   420
  by (erule arg_cong)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   421
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   422
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
   423
  by (simp add: split_eta)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   424
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   425
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   426
  Simplification procedure for @{thm [source] cond_split_eta}.  Using
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   427
  @{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
   428
  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
   429
  existing proofs very inefficient; similarly for @{text
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   430
  split_beta}.
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   431
*}
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   432
26480
544cef16045b replaced 'ML_setup' by 'ML';
wenzelm
parents: 26358
diff changeset
   433
ML {*
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   434
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   435
local
18328
841261f303a1 simprocs: static evaluation of simpset;
wenzelm
parents: 18220
diff changeset
   436
  val cond_split_eta_ss = HOL_basic_ss addsimps [thm "cond_split_eta"]
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   437
  fun  Pair_pat k 0 (Bound m) = (m = k)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   438
  |    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
   439
                        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
   440
  |    Pair_pat _ _ _ = false;
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   441
  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
   442
  |   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
   443
  |   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
   444
  |   no_args _ _ _ = true;
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   445
  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
   446
  |   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
   447
  |   split_pat tp i _ = NONE;
20044
92cc2f4c7335 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19656
diff changeset
   448
  fun metaeq ss lhs rhs = mk_meta_eq (Goal.prove (Simplifier.the_context ss) [] []
13480
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13462
diff changeset
   449
        (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs,rhs)))
18328
841261f303a1 simprocs: static evaluation of simpset;
wenzelm
parents: 18220
diff changeset
   450
        (K (simp_tac (Simplifier.inherit_context ss cond_split_eta_ss) 1)));
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   451
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   452
  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
   453
  |   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
   454
                        (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
   455
  |   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
   456
  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
   457
  |    eta_term_pat _ _ _ = false;
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   458
  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
   459
  |   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
   460
                              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
   461
  |   subst arg k i t = t;
20044
92cc2f4c7335 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19656
diff changeset
   462
  fun beta_proc ss (s as Const ("split", _) $ Abs (_, _, t) $ arg) =
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   463
        (case split_pat beta_term_pat 1 t of
20044
92cc2f4c7335 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19656
diff changeset
   464
        SOME (i,f) => SOME (metaeq ss s (subst arg 0 i f))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   465
        | NONE => NONE)
20044
92cc2f4c7335 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19656
diff changeset
   466
  |   beta_proc _ _ = NONE;
92cc2f4c7335 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19656
diff changeset
   467
  fun eta_proc ss (s as Const ("split", _) $ Abs (_, _, t)) =
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   468
        (case split_pat eta_term_pat 1 t of
20044
92cc2f4c7335 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19656
diff changeset
   469
          SOME (_,ft) => SOME (metaeq ss s (let val (f $ arg) = ft in f end))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   470
        | NONE => NONE)
20044
92cc2f4c7335 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19656
diff changeset
   471
  |   eta_proc _ _ = NONE;
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   472
in
32010
cb1a1c94b4cd more antiquotations;
wenzelm
parents: 31775
diff changeset
   473
  val split_beta_proc = Simplifier.simproc @{theory} "split_beta" ["split f z"] (K beta_proc);
cb1a1c94b4cd more antiquotations;
wenzelm
parents: 31775
diff changeset
   474
  val split_eta_proc = Simplifier.simproc @{theory} "split_eta" ["split f"] (K eta_proc);
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   475
end;
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   476
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   477
Addsimprocs [split_beta_proc, split_eta_proc];
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
26798
a9134a089106 split_beta is now declared as monotonicity rule, to allow bounded
berghofe
parents: 26588
diff changeset
   480
lemma split_beta [mono]: "(%(x, y). P x y) z = P (fst z) (snd z)"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   481
  by (subst surjective_pairing, rule split_conv)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   482
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24162
diff changeset
   483
lemma split_split [noatp]: "R(split c p) = (ALL x y. p = (x, y) --> R(c x y))"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   484
  -- {* For use with @{text split} and the Simplifier. *}
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15422
diff changeset
   485
  by (insert surj_pair [of p], clarify, simp)
11838
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
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   488
  @{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
   489
  done after the Splitter has been speeded up significantly;
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   490
  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
   491
  current goal contains one of those constants.
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   492
*}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   493
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24162
diff changeset
   494
lemma split_split_asm [noatp]: "R (split c p) = (~(EX x y. p = (x, y) & (~R (c x y))))"
14208
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   495
by (subst split_split, simp)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   496
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   497
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   498
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   499
  \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
   500
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   501
  \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
   502
  call @{text simp} using @{thm [source] split} as rewrite. *}
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 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
   505
  apply (simp only: split_tupled_all)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   506
  apply (simp (no_asm_simp))
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   507
  done
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   508
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   509
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
   510
  apply (simp only: split_tupled_all)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   511
  apply (simp (no_asm_simp))
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   512
  done
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   513
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   514
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
   515
  by (induct p) (auto simp add: split_def)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   516
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   517
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
   518
  by (induct p) (auto simp add: split_def)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   519
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   520
lemma splitE2:
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   521
  "[| 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
   522
proof -
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   523
  assume q: "Q (split P z)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   524
  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
   525
  show R
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   526
    apply (rule r surjective_pairing)+
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   527
    apply (rule split_beta [THEN subst], rule q)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   528
    done
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   529
qed
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   530
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   531
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
   532
  by simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   533
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   534
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
   535
  by simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   536
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   537
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
   538
by (simp only: split_tupled_all, simp)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   539
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   540
lemma mem_splitE:
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   541
  assumes major: "z: split c p"
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   542
    and cases: "!!x y. [| p = (x,y); z: c x y |] ==> Q"
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   543
  shows Q
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   544
  by (rule major [unfolded split_def] cases surjective_pairing)+
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   545
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   546
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
   547
declare mem_splitE [elim!] splitE' [elim!] splitE [elim!]
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   548
26340
a85fe32e7b2f more antiquotations;
wenzelm
parents: 26143
diff changeset
   549
ML {*
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   550
local (* filtering with exists_p_split is an essential optimization *)
16121
wenzelm
parents: 15570
diff changeset
   551
  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
   552
    | 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
   553
    | exists_p_split (Abs (_, _, t)) = exists_p_split t
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   554
    | exists_p_split _ = false;
16121
wenzelm
parents: 15570
diff changeset
   555
  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
   556
in
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   557
val split_conv_tac = SUBGOAL (fn (t, i) =>
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   558
    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
   559
end;
26340
a85fe32e7b2f more antiquotations;
wenzelm
parents: 26143
diff changeset
   560
*}
a85fe32e7b2f more antiquotations;
wenzelm
parents: 26143
diff changeset
   561
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   562
(* 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
   563
   to quite time-consuming computations (in particular for nested tuples) *)
26340
a85fe32e7b2f more antiquotations;
wenzelm
parents: 26143
diff changeset
   564
declaration {* fn _ =>
a85fe32e7b2f more antiquotations;
wenzelm
parents: 26143
diff changeset
   565
  Classical.map_cs (fn cs => cs addSbefore ("split_conv_tac", split_conv_tac))
16121
wenzelm
parents: 15570
diff changeset
   566
*}
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   567
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24162
diff changeset
   568
lemma split_eta_SetCompr [simp,noatp]: "(%u. EX x y. u = (x, y) & P (x, y)) = P"
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   569
  by (rule ext) fast
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   570
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24162
diff changeset
   571
lemma split_eta_SetCompr2 [simp,noatp]: "(%u. EX x y. u = (x, y) & P x y) = split P"
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   572
  by (rule ext) fast
11838
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
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
   575
  -- {* Allows simplifications of nested splits in case of independent predicates. *}
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   576
  by (rule ext) blast
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   577
14337
e13731554e50 undid split_comp_eq[simp] because it leads to nontermination together with split_def!
nipkow
parents: 14208
diff changeset
   578
(* 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
   579
   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
   580
   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
   581
*)
e13731554e50 undid split_comp_eq[simp] because it leads to nontermination together with split_def!
nipkow
parents: 14208
diff changeset
   582
lemma split_comp_eq: 
20415
e3d2d7b01279 explicit type variables prevent empty sorts
paulson
parents: 20380
diff changeset
   583
  fixes f :: "'a => 'b => 'c" and g :: "'d => 'a"
e3d2d7b01279 explicit type variables prevent empty sorts
paulson
parents: 20380
diff changeset
   584
  shows "(%u. f (g (fst u)) (snd u)) = (split (%x. f (g x)))"
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   585
  by (rule ext) auto
14101
d25c23e46173 added upd_fst, upd_snd, some thms
oheimb
parents: 13480
diff changeset
   586
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   587
lemma pair_imageI [intro]: "(a, b) : A ==> f a b : (%(a, b). f a b) ` A"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   588
  apply (rule_tac x = "(a, b)" in image_eqI)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   589
   apply auto
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   590
  done
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   591
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   592
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
   593
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   594
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   595
(*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   596
the following  would be slightly more general,
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   597
but cannot be used as rewrite rule:
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   598
### 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
   599
### ?y = .x
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   600
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
   601
by (rtac some_equality 1)
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   602
by ( Simp_tac 1)
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   603
by (split_all_tac 1)
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   604
by (Asm_full_simp_tac 1)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   605
qed "The_split_eq";
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   606
*)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   607
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   608
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   609
  Setup of internal @{text split_rule}.
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   610
*}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   611
25511
54db9b5080b8 more canonical attribute application
haftmann
parents: 24844
diff changeset
   612
definition
54db9b5080b8 more canonical attribute application
haftmann
parents: 24844
diff changeset
   613
  internal_split :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c"
54db9b5080b8 more canonical attribute application
haftmann
parents: 24844
diff changeset
   614
where
11032
83f723e86dac added hidden internal_split constant;
wenzelm
parents: 11025
diff changeset
   615
  "internal_split == split"
83f723e86dac added hidden internal_split constant;
wenzelm
parents: 11025
diff changeset
   616
83f723e86dac added hidden internal_split constant;
wenzelm
parents: 11025
diff changeset
   617
lemma internal_split_conv: "internal_split c (a, b) = c a b"
83f723e86dac added hidden internal_split constant;
wenzelm
parents: 11025
diff changeset
   618
  by (simp only: internal_split_def split_conv)
83f723e86dac added hidden internal_split constant;
wenzelm
parents: 11025
diff changeset
   619
83f723e86dac added hidden internal_split constant;
wenzelm
parents: 11025
diff changeset
   620
hide const internal_split
83f723e86dac added hidden internal_split constant;
wenzelm
parents: 11025
diff changeset
   621
11025
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   622
use "Tools/split_rule.ML"
11032
83f723e86dac added hidden internal_split constant;
wenzelm
parents: 11025
diff changeset
   623
setup SplitRule.setup
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   624
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   625
lemmas prod_caseI = prod.cases [THEN iffD2, standard]
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   626
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   627
lemma prod_caseI2: "!!p. [| !!a b. p = (a, b) ==> c a b |] ==> prod_case c p"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   628
  by auto
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   629
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   630
lemma prod_caseI2': "!!p. [| !!a b. (a, b) = p ==> c a b x |] ==> prod_case c p x"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   631
  by (auto simp: split_tupled_all)
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   632
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   633
lemma prod_caseE: "prod_case c p ==> (!!x y. p = (x, y) ==> c x y ==> Q) ==> Q"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   634
  by (induct p) auto
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   635
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   636
lemma prod_caseE': "prod_case c p z ==> (!!x y. p = (x, y) ==> c x y z ==> Q) ==> Q"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   637
  by (induct p) auto
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   638
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   639
lemma prod_case_unfold: "prod_case = (%c p. c (fst p) (snd p))"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   640
  by (simp add: expand_fun_eq)
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   641
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   642
declare prod_caseI2' [intro!] prod_caseI2 [intro!] prod_caseI [intro!]
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   643
declare prod_caseE' [elim!] prod_caseE [elim!]
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   644
24844
98c006a30218 certificates for code generator case expressions
haftmann
parents: 24699
diff changeset
   645
lemma prod_case_split:
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   646
  "prod_case = split"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   647
  by (auto simp add: expand_fun_eq)
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   648
26143
314c0bcb7df7 Added useful general lemmas from the work with the HeapMonad
bulwahn
parents: 25885
diff changeset
   649
lemma prod_case_beta:
314c0bcb7df7 Added useful general lemmas from the work with the HeapMonad
bulwahn
parents: 25885
diff changeset
   650
  "prod_case f p = f (fst p) (snd p)"
314c0bcb7df7 Added useful general lemmas from the work with the HeapMonad
bulwahn
parents: 25885
diff changeset
   651
  unfolding prod_case_split split_beta ..
314c0bcb7df7 Added useful general lemmas from the work with the HeapMonad
bulwahn
parents: 25885
diff changeset
   652
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   653
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   654
subsection {* Further cases/induct rules for tuples *}
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   655
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   656
lemma prod_cases3 [cases type]:
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   657
  obtains (fields) a b c where "y = (a, b, c)"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   658
  by (cases y, case_tac b) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   659
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   660
lemma prod_induct3 [case_names fields, induct type]:
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   661
    "(!!a b c. P (a, b, c)) ==> P x"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   662
  by (cases x) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   663
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   664
lemma prod_cases4 [cases type]:
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   665
  obtains (fields) a b c d where "y = (a, b, c, d)"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   666
  by (cases y, case_tac c) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   667
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   668
lemma prod_induct4 [case_names fields, induct type]:
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   669
    "(!!a b c d. P (a, b, c, d)) ==> P x"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   670
  by (cases x) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   671
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   672
lemma prod_cases5 [cases type]:
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   673
  obtains (fields) a b c d e where "y = (a, b, c, d, e)"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   674
  by (cases y, case_tac d) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   675
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   676
lemma prod_induct5 [case_names fields, induct type]:
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   677
    "(!!a b c d e. P (a, b, c, d, e)) ==> P x"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   678
  by (cases x) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   679
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   680
lemma prod_cases6 [cases type]:
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   681
  obtains (fields) a b c d e f where "y = (a, b, c, d, e, f)"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   682
  by (cases y, case_tac e) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   683
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   684
lemma prod_induct6 [case_names fields, induct type]:
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   685
    "(!!a b c d e f. P (a, b, c, d, e, f)) ==> P x"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   686
  by (cases x) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   687
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   688
lemma prod_cases7 [cases type]:
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   689
  obtains (fields) a b c d e f g where "y = (a, b, c, d, e, f, g)"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   690
  by (cases y, case_tac f) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   691
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   692
lemma prod_induct7 [case_names fields, induct type]:
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   693
    "(!!a b c d e f g. P (a, b, c, d, e, f, g)) ==> P x"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   694
  by (cases x) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   695
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   696
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   697
subsubsection {* Derived operations *}
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   698
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   699
text {*
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   700
  The composition-uncurry combinator.
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   701
*}
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   702
26588
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26480
diff changeset
   703
notation fcomp (infixl "o>" 60)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   704
26588
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26480
diff changeset
   705
definition
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26480
diff changeset
   706
  scomp :: "('a \<Rightarrow> 'b \<times> 'c) \<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> 'd) \<Rightarrow> 'a \<Rightarrow> 'd" (infixl "o\<rightarrow>" 60)
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26480
diff changeset
   707
where
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26480
diff changeset
   708
  "f o\<rightarrow> g = (\<lambda>x. split g (f x))"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   709
26588
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26480
diff changeset
   710
lemma scomp_apply:  "(f o\<rightarrow> g) x = split g (f x)"
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26480
diff changeset
   711
  by (simp add: scomp_def)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   712
26588
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26480
diff changeset
   713
lemma Pair_scomp: "Pair x o\<rightarrow> f = f x"
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26480
diff changeset
   714
  by (simp add: expand_fun_eq scomp_apply)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   715
26588
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26480
diff changeset
   716
lemma scomp_Pair: "x o\<rightarrow> Pair = x"
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26480
diff changeset
   717
  by (simp add: expand_fun_eq scomp_apply)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   718
26588
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26480
diff changeset
   719
lemma scomp_scomp: "(f o\<rightarrow> g) o\<rightarrow> h = f o\<rightarrow> (\<lambda>x. g x o\<rightarrow> h)"
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26480
diff changeset
   720
  by (simp add: expand_fun_eq split_twice scomp_def)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   721
26588
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26480
diff changeset
   722
lemma scomp_fcomp: "(f o\<rightarrow> g) o> h = f o\<rightarrow> (\<lambda>x. g x o> h)"
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26480
diff changeset
   723
  by (simp add: expand_fun_eq scomp_apply fcomp_def split_def)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   724
26588
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26480
diff changeset
   725
lemma fcomp_scomp: "(f o> g) o\<rightarrow> h = f o> (g o\<rightarrow> h)"
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26480
diff changeset
   726
  by (simp add: expand_fun_eq scomp_apply fcomp_apply)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   727
31202
52d332f8f909 pretty printing of functional combinators for evaluation code
haftmann
parents: 30924
diff changeset
   728
code_const scomp
52d332f8f909 pretty printing of functional combinators for evaluation code
haftmann
parents: 30924
diff changeset
   729
  (Eval infixl 3 "#->")
52d332f8f909 pretty printing of functional combinators for evaluation code
haftmann
parents: 30924
diff changeset
   730
26588
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26480
diff changeset
   731
no_notation fcomp (infixl "o>" 60)
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26480
diff changeset
   732
no_notation scomp (infixl "o\<rightarrow>" 60)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   733
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   734
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   735
text {*
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   736
  @{term prod_fun} --- action of the product functor upon
31775
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 31723
diff changeset
   737
  Datatypes.
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   738
*}
21195
0cca8d19557d two further lemmas on split
haftmann
parents: 21046
diff changeset
   739
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   740
definition prod_fun :: "('a \<Rightarrow> 'c) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c \<times> 'd" where
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28537
diff changeset
   741
  [code del]: "prod_fun f g = (\<lambda>(x, y). (f x, g y))"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   742
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28537
diff changeset
   743
lemma prod_fun [simp, code]: "prod_fun f g (a, b) = (f a, g b)"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   744
  by (simp add: prod_fun_def)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   745
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   746
lemma prod_fun_compose: "prod_fun (f1 o f2) (g1 o g2) = (prod_fun f1 g1 o prod_fun f2 g2)"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   747
  by (rule ext) auto
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   748
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   749
lemma prod_fun_ident [simp]: "prod_fun (%x. x) (%y. y) = (%z. z)"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   750
  by (rule ext) auto
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   751
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   752
lemma prod_fun_imageI [intro]: "(a, b) : r ==> (f a, g b) : prod_fun f g ` r"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   753
  apply (rule image_eqI)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   754
  apply (rule prod_fun [symmetric], assumption)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   755
  done
21195
0cca8d19557d two further lemmas on split
haftmann
parents: 21046
diff changeset
   756
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   757
lemma prod_fun_imageE [elim!]:
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   758
  assumes major: "c: (prod_fun f g)`r"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   759
    and cases: "!!x y. [| c=(f(x),g(y));  (x,y):r |] ==> P"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   760
  shows P
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   761
  apply (rule major [THEN imageE])
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   762
  apply (rule_tac p = x in PairE)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   763
  apply (rule cases)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   764
   apply (blast intro: prod_fun)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   765
  apply blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   766
  done
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   767
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   768
definition
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   769
  apfst :: "('a \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c \<times> 'b"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   770
where
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28537
diff changeset
   771
  [code del]: "apfst f = prod_fun f id"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   772
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   773
definition
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   774
  apsnd :: "('b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'a \<times> 'c"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   775
where
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28537
diff changeset
   776
  [code del]: "apsnd f = prod_fun id f"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   777
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   778
lemma apfst_conv [simp, code]:
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   779
  "apfst f (x, y) = (f x, y)" 
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   780
  by (simp add: apfst_def)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   781
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   782
lemma upd_snd_conv [simp, code]:
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   783
  "apsnd f (x, y) = (x, f y)" 
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   784
  by (simp add: apsnd_def)
21195
0cca8d19557d two further lemmas on split
haftmann
parents: 21046
diff changeset
   785
0cca8d19557d two further lemmas on split
haftmann
parents: 21046
diff changeset
   786
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   787
text {*
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   788
  Disjoint union of a family of sets -- Sigma.
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   789
*}
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   790
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   791
definition  Sigma :: "['a set, 'a => 'b set] => ('a \<times> 'b) set" where
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   792
  Sigma_def: "Sigma A B == UN x:A. UN y:B x. {Pair x y}"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   793
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   794
abbreviation
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   795
  Times :: "['a set, 'b set] => ('a * 'b) set"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   796
    (infixr "<*>" 80) where
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   797
  "A <*> B == Sigma A (%_. B)"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   798
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   799
notation (xsymbols)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   800
  Times  (infixr "\<times>" 80)
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   801
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   802
notation (HTML output)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   803
  Times  (infixr "\<times>" 80)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   804
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   805
syntax
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   806
  "@Sigma" ::"[pttrn, 'a set, 'b set] => ('a * 'b) set" ("(3SIGMA _:_./ _)" [0, 0, 10] 10)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   807
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   808
translations
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   809
  "SIGMA x:A. B" == "Product_Type.Sigma A (%x. B)"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   810
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   811
lemma SigmaI [intro!]: "[| a:A;  b:B(a) |] ==> (a,b) : Sigma A B"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   812
  by (unfold Sigma_def) blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   813
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   814
lemma SigmaE [elim!]:
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   815
    "[| c: Sigma A B;
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   816
        !!x y.[| x:A;  y:B(x);  c=(x,y) |] ==> P
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   817
     |] ==> P"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   818
  -- {* The general elimination rule. *}
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   819
  by (unfold Sigma_def) blast
20588
c847c56edf0c added operational equality
haftmann
parents: 20415
diff changeset
   820
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   821
text {*
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   822
  Elimination of @{term "(a, b) : A \<times> B"} -- introduces no
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   823
  eigenvariables.
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   824
*}
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   825
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   826
lemma SigmaD1: "(a, b) : Sigma A B ==> a : A"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   827
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   828
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   829
lemma SigmaD2: "(a, b) : Sigma A B ==> b : B a"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   830
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   831
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   832
lemma SigmaE2:
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   833
    "[| (a, b) : Sigma A B;
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   834
        [| a:A;  b:B(a) |] ==> P
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   835
     |] ==> P"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   836
  by blast
20588
c847c56edf0c added operational equality
haftmann
parents: 20415
diff changeset
   837
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   838
lemma Sigma_cong:
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   839
     "\<lbrakk>A = B; !!x. x \<in> B \<Longrightarrow> C x = D x\<rbrakk>
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   840
      \<Longrightarrow> (SIGMA x: A. C x) = (SIGMA x: B. D x)"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   841
  by auto
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   842
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   843
lemma Sigma_mono: "[| A <= C; !!x. x:A ==> B x <= D x |] ==> Sigma A B <= Sigma C D"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   844
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   845
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   846
lemma Sigma_empty1 [simp]: "Sigma {} B = {}"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   847
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   848
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   849
lemma Sigma_empty2 [simp]: "A <*> {} = {}"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   850
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   851
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   852
lemma UNIV_Times_UNIV [simp]: "UNIV <*> UNIV = UNIV"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   853
  by auto
21908
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
   854
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   855
lemma Compl_Times_UNIV1 [simp]: "- (UNIV <*> A) = UNIV <*> (-A)"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   856
  by auto
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   857
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   858
lemma Compl_Times_UNIV2 [simp]: "- (A <*> UNIV) = (-A) <*> UNIV"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   859
  by auto
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   860
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   861
lemma mem_Sigma_iff [iff]: "((a,b): Sigma A B) = (a:A & b:B(a))"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   862
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   863
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   864
lemma Times_subset_cancel2: "x:C ==> (A <*> C <= B <*> C) = (A <= B)"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   865
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   866
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   867
lemma Times_eq_cancel2: "x:C ==> (A <*> C = B <*> C) = (A = B)"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   868
  by (blast elim: equalityE)
20588
c847c56edf0c added operational equality
haftmann
parents: 20415
diff changeset
   869
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   870
lemma SetCompr_Sigma_eq:
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   871
    "Collect (split (%x y. P x & Q x y)) = (SIGMA x:Collect P. Collect (Q x))"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   872
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   873
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   874
lemma Collect_split [simp]: "{(a,b). P a & Q b} = Collect P <*> Collect Q"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   875
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   876
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   877
lemma UN_Times_distrib:
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   878
  "(UN (a,b):(A <*> B). E a <*> F b) = (UNION A E) <*> (UNION B F)"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   879
  -- {* Suggested by Pierre Chartier *}
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   880
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   881
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   882
lemma split_paired_Ball_Sigma [simp,noatp]:
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   883
    "(ALL z: Sigma A B. P z) = (ALL x:A. ALL y: B x. P(x,y))"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   884
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   885
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   886
lemma split_paired_Bex_Sigma [simp,noatp]:
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   887
    "(EX z: Sigma A B. P z) = (EX x:A. EX y: B x. P(x,y))"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   888
  by blast
21908
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
   889
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   890
lemma Sigma_Un_distrib1: "(SIGMA i:I Un J. C(i)) = (SIGMA i:I. C(i)) Un (SIGMA j:J. C(j))"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   891
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   892
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   893
lemma Sigma_Un_distrib2: "(SIGMA i:I. A(i) Un B(i)) = (SIGMA i:I. A(i)) Un (SIGMA i:I. B(i))"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   894
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   895
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   896
lemma Sigma_Int_distrib1: "(SIGMA i:I Int J. C(i)) = (SIGMA i:I. C(i)) Int (SIGMA j:J. C(j))"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   897
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   898
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   899
lemma Sigma_Int_distrib2: "(SIGMA i:I. A(i) Int B(i)) = (SIGMA i:I. A(i)) Int (SIGMA i:I. B(i))"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   900
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   901
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   902
lemma Sigma_Diff_distrib1: "(SIGMA i:I - J. C(i)) = (SIGMA i:I. C(i)) - (SIGMA j:J. C(j))"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   903
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   904
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   905
lemma Sigma_Diff_distrib2: "(SIGMA i:I. A(i) - B(i)) = (SIGMA i:I. A(i)) - (SIGMA i:I. B(i))"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   906
  by blast
21908
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
   907
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   908
lemma Sigma_Union: "Sigma (Union X) B = (UN A:X. Sigma A B)"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   909
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   910
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   911
text {*
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   912
  Non-dependent versions are needed to avoid the need for higher-order
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   913
  matching, especially when the rules are re-oriented.
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   914
*}
21908
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
   915
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   916
lemma Times_Un_distrib1: "(A Un B) <*> C = (A <*> C) Un (B <*> C)"
28719
01e04e41cc7b added lemma
nipkow
parents: 28562
diff changeset
   917
by blast
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   918
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   919
lemma Times_Int_distrib1: "(A Int B) <*> C = (A <*> C) Int (B <*> C)"
28719
01e04e41cc7b added lemma
nipkow
parents: 28562
diff changeset
   920
by blast
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   921
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   922
lemma Times_Diff_distrib1: "(A - B) <*> C = (A <*> C) - (B <*> C)"
28719
01e04e41cc7b added lemma
nipkow
parents: 28562
diff changeset
   923
by blast
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   924
28719
01e04e41cc7b added lemma
nipkow
parents: 28562
diff changeset
   925
lemma insert_times_insert[simp]:
01e04e41cc7b added lemma
nipkow
parents: 28562
diff changeset
   926
  "insert a A \<times> insert b B =
01e04e41cc7b added lemma
nipkow
parents: 28562
diff changeset
   927
   insert (a,b) (A \<times> insert b B \<union> insert a A \<times> B)"
01e04e41cc7b added lemma
nipkow
parents: 28562
diff changeset
   928
by blast
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   929
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   930
subsubsection {* Code generator setup *}
21908
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
   931
20588
c847c56edf0c added operational equality
haftmann
parents: 20415
diff changeset
   932
instance * :: (eq, eq) eq ..
c847c56edf0c added operational equality
haftmann
parents: 20415
diff changeset
   933
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28537
diff changeset
   934
lemma [code]:
28346
b8390cd56b8f discontinued special treatment of op = vs. eq_class.eq
haftmann
parents: 28262
diff changeset
   935
  "eq_class.eq (x1\<Colon>'a\<Colon>eq, y1\<Colon>'b\<Colon>eq) (x2, y2) \<longleftrightarrow> x1 = x2 \<and> y1 = y2" by (simp add: eq)
20588
c847c56edf0c added operational equality
haftmann
parents: 20415
diff changeset
   936
24844
98c006a30218 certificates for code generator case expressions
haftmann
parents: 24699
diff changeset
   937
lemma split_case_cert:
98c006a30218 certificates for code generator case expressions
haftmann
parents: 24699
diff changeset
   938
  assumes "CASE \<equiv> split f"
98c006a30218 certificates for code generator case expressions
haftmann
parents: 24699
diff changeset
   939
  shows "CASE (a, b) \<equiv> f a b"
98c006a30218 certificates for code generator case expressions
haftmann
parents: 24699
diff changeset
   940
  using assms by simp
98c006a30218 certificates for code generator case expressions
haftmann
parents: 24699
diff changeset
   941
98c006a30218 certificates for code generator case expressions
haftmann
parents: 24699
diff changeset
   942
setup {*
98c006a30218 certificates for code generator case expressions
haftmann
parents: 24699
diff changeset
   943
  Code.add_case @{thm split_case_cert}
98c006a30218 certificates for code generator case expressions
haftmann
parents: 24699
diff changeset
   944
*}
98c006a30218 certificates for code generator case expressions
haftmann
parents: 24699
diff changeset
   945
21908
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
   946
code_type *
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
   947
  (SML infix 2 "*")
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
   948
  (OCaml infix 2 "*")
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
   949
  (Haskell "!((_),/ (_))")
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
   950
20588
c847c56edf0c added operational equality
haftmann
parents: 20415
diff changeset
   951
code_instance * :: eq
c847c56edf0c added operational equality
haftmann
parents: 20415
diff changeset
   952
  (Haskell -)
c847c56edf0c added operational equality
haftmann
parents: 20415
diff changeset
   953
28346
b8390cd56b8f discontinued special treatment of op = vs. eq_class.eq
haftmann
parents: 28262
diff changeset
   954
code_const "eq_class.eq \<Colon> 'a\<Colon>eq \<times> 'b\<Colon>eq \<Rightarrow> 'a \<times> 'b \<Rightarrow> bool"
20588
c847c56edf0c added operational equality
haftmann
parents: 20415
diff changeset
   955
  (Haskell infixl 4 "==")
c847c56edf0c added operational equality
haftmann
parents: 20415
diff changeset
   956
21908
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
   957
code_const Pair
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
   958
  (SML "!((_),/ (_))")
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
   959
  (OCaml "!((_),/ (_))")
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
   960
  (Haskell "!((_),/ (_))")
20588
c847c56edf0c added operational equality
haftmann
parents: 20415
diff changeset
   961
22389
bdf16741d039 using "fst" "snd" for Haskell code
haftmann
parents: 22349
diff changeset
   962
code_const fst and snd
bdf16741d039 using "fst" "snd" for Haskell code
haftmann
parents: 22349
diff changeset
   963
  (Haskell "fst" and "snd")
bdf16741d039 using "fst" "snd" for Haskell code
haftmann
parents: 22349
diff changeset
   964
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   965
types_code
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   966
  "*"     ("(_ */ _)")
16770
1f1b1fae30e4 Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents: 16634
diff changeset
   967
attach (term_of) {*
25885
6fbc3f54f819 New interface for test data generators.
berghofe
parents: 25534
diff changeset
   968
fun term_of_id_42 aF aT bF bT (x, y) = HOLogic.pair_const aT bT $ aF x $ bF y;
16770
1f1b1fae30e4 Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents: 16634
diff changeset
   969
*}
1f1b1fae30e4 Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents: 16634
diff changeset
   970
attach (test) {*
25885
6fbc3f54f819 New interface for test data generators.
berghofe
parents: 25534
diff changeset
   971
fun gen_id_42 aG aT bG bT i =
6fbc3f54f819 New interface for test data generators.
berghofe
parents: 25534
diff changeset
   972
  let
6fbc3f54f819 New interface for test data generators.
berghofe
parents: 25534
diff changeset
   973
    val (x, t) = aG i;
6fbc3f54f819 New interface for test data generators.
berghofe
parents: 25534
diff changeset
   974
    val (y, u) = bG i
6fbc3f54f819 New interface for test data generators.
berghofe
parents: 25534
diff changeset
   975
  in ((x, y), fn () => HOLogic.pair_const aT bT $ t () $ u ()) end;
16770
1f1b1fae30e4 Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents: 16634
diff changeset
   976
*}
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   977
18706
1e7562c7afe6 Re-inserted consts_code declaration accidentally deleted
berghofe
parents: 18702
diff changeset
   978
consts_code
1e7562c7afe6 Re-inserted consts_code declaration accidentally deleted
berghofe
parents: 18702
diff changeset
   979
  "Pair"    ("(_,/ _)")
1e7562c7afe6 Re-inserted consts_code declaration accidentally deleted
berghofe
parents: 18702
diff changeset
   980
21908
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
   981
setup {*
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
   982
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
   983
let
18013
3f5d0acdfdba added extraction interface for code generator
haftmann
parents: 17956
diff changeset
   984
19039
8eae46249628 added theory of executable rational numbers
haftmann
parents: 19008
diff changeset
   985
fun strip_abs_split 0 t = ([], t)
8eae46249628 added theory of executable rational numbers
haftmann
parents: 19008
diff changeset
   986
  | strip_abs_split i (Abs (s, T, t)) =
18013
3f5d0acdfdba added extraction interface for code generator
haftmann
parents: 17956
diff changeset
   987
      let
3f5d0acdfdba added extraction interface for code generator
haftmann
parents: 17956
diff changeset
   988
        val s' = Codegen.new_name t s;
3f5d0acdfdba added extraction interface for code generator
haftmann
parents: 17956
diff changeset
   989
        val v = Free (s', T)
19039
8eae46249628 added theory of executable rational numbers
haftmann
parents: 19008
diff changeset
   990
      in apfst (cons v) (strip_abs_split (i-1) (subst_bound (v, t))) end
8eae46249628 added theory of executable rational numbers
haftmann
parents: 19008
diff changeset
   991
  | strip_abs_split i (u as Const ("split", _) $ t) = (case strip_abs_split (i+1) t of
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   992
        (v :: v' :: vs, u) => (HOLogic.mk_prod (v, v') :: vs, u)
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   993
      | _ => ([], u))
30604
2a9911f4b0a3 split_codegen now eta-expands terms on-the-fly.
berghofe
parents: 28719
diff changeset
   994
  | strip_abs_split i t =
2a9911f4b0a3 split_codegen now eta-expands terms on-the-fly.
berghofe
parents: 28719
diff changeset
   995
      strip_abs_split i (Abs ("x", hd (binder_types (fastype_of t)), t $ Bound 0));
18013
3f5d0acdfdba added extraction interface for code generator
haftmann
parents: 17956
diff changeset
   996
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28346
diff changeset
   997
fun let_codegen thy defs dep thyname brack t gr = (case strip_comb t of
16634
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
   998
    (t1 as Const ("Let", _), t2 :: t3 :: ts) =>
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   999
    let
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
  1000
      fun dest_let (l as Const ("Let", _) $ t $ u) =
19039
8eae46249628 added theory of executable rational numbers
haftmann
parents: 19008
diff changeset
  1001
          (case strip_abs_split 1 u of
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
  1002
             ([p], u') => apfst (cons (p, t)) (dest_let u')
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
  1003
           | _ => ([], l))
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
  1004
        | dest_let t = ([], t);
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28346
diff changeset
  1005
      fun mk_code (l, r) gr =
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
  1006
        let
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28346
diff changeset
  1007
          val (pl, gr1) = Codegen.invoke_codegen thy defs dep thyname false l gr;
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28346
diff changeset
  1008
          val (pr, gr2) = Codegen.invoke_codegen thy defs dep thyname false r gr1;
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28346
diff changeset
  1009
        in ((pl, pr), gr2) end
16634
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
  1010
    in case dest_let (t1 $ t2 $ t3) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
  1011
        ([], _) => NONE
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
  1012
      | (ps, u) =>
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
  1013
          let
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28346
diff changeset
  1014
            val (qs, gr1) = fold_map mk_code ps gr;
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28346
diff changeset
  1015
            val (pu, gr2) = Codegen.invoke_codegen thy defs dep thyname false u gr1;
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28346
diff changeset
  1016
            val (pargs, gr3) = fold_map
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28346
diff changeset
  1017
              (Codegen.invoke_codegen thy defs dep thyname true) ts gr2
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
  1018
          in
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28346
diff changeset
  1019
            SOME (Codegen.mk_app brack
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32010
diff changeset
  1020
              (Pretty.blk (0, [Codegen.str "let ", Pretty.blk (0, flat
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26798
diff changeset
  1021
                  (separate [Codegen.str ";", Pretty.brk 1] (map (fn (pl, pr) =>
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26798
diff changeset
  1022
                    [Pretty.block [Codegen.str "val ", pl, Codegen.str " =",
16634
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
  1023
                       Pretty.brk 1, pr]]) qs))),
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26798
diff changeset
  1024
                Pretty.brk 1, Codegen.str "in ", pu,
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28346
diff changeset
  1025
                Pretty.brk 1, Codegen.str "end"])) pargs, gr3)
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
  1026
          end
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
  1027
    end
16634
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
  1028
  | _ => NONE);
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
  1029
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28346
diff changeset
  1030
fun split_codegen thy defs dep thyname brack t gr = (case strip_comb t of
16634
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
  1031
    (t1 as Const ("split", _), t2 :: ts) =>
30604
2a9911f4b0a3 split_codegen now eta-expands terms on-the-fly.
berghofe
parents: 28719
diff changeset
  1032
      let
2a9911f4b0a3 split_codegen now eta-expands terms on-the-fly.
berghofe
parents: 28719
diff changeset
  1033
        val ([p], u) = strip_abs_split 1 (t1 $ t2);
2a9911f4b0a3 split_codegen now eta-expands terms on-the-fly.
berghofe
parents: 28719
diff changeset
  1034
        val (q, gr1) = Codegen.invoke_codegen thy defs dep thyname false p gr;
2a9911f4b0a3 split_codegen now eta-expands terms on-the-fly.
berghofe
parents: 28719
diff changeset
  1035
        val (pu, gr2) = Codegen.invoke_codegen thy defs dep thyname false u gr1;
2a9911f4b0a3 split_codegen now eta-expands terms on-the-fly.
berghofe
parents: 28719
diff changeset
  1036
        val (pargs, gr3) = fold_map
2a9911f4b0a3 split_codegen now eta-expands terms on-the-fly.
berghofe
parents: 28719
diff changeset
  1037
          (Codegen.invoke_codegen thy defs dep thyname true) ts gr2
2a9911f4b0a3 split_codegen now eta-expands terms on-the-fly.
berghofe
parents: 28719
diff changeset
  1038
      in
2a9911f4b0a3 split_codegen now eta-expands terms on-the-fly.
berghofe
parents: 28719
diff changeset
  1039
        SOME (Codegen.mk_app brack
2a9911f4b0a3 split_codegen now eta-expands terms on-the-fly.
berghofe
parents: 28719
diff changeset
  1040
          (Pretty.block [Codegen.str "(fn ", q, Codegen.str " =>",
2a9911f4b0a3 split_codegen now eta-expands terms on-the-fly.
berghofe
parents: 28719
diff changeset
  1041
            Pretty.brk 1, pu, Codegen.str ")"]) pargs, gr2)
2a9911f4b0a3 split_codegen now eta-expands terms on-the-fly.
berghofe
parents: 28719
diff changeset
  1042
      end
16634
f19d58cfb47a Adapted to new interface of code generator.
berghofe
parents: 16417
diff changeset
  1043
  | _ => NONE);
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
  1044
21908
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
  1045
in
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
  1046
20105
454f4be984b7 adaptions in codegen
haftmann
parents: 20044
diff changeset
  1047
  Codegen.add_codegen "let_codegen" let_codegen
454f4be984b7 adaptions in codegen
haftmann
parents: 20044
diff changeset
  1048
  #> Codegen.add_codegen "split_codegen" split_codegen
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
  1049
21908
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
  1050
end
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
  1051
*}
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
  1052
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
  1053
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
  1054
subsection {* Legacy bindings *}
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
  1055
21908
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
  1056
ML {*
15404
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1057
val Collect_split = thm "Collect_split";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1058
val Compl_Times_UNIV1 = thm "Compl_Times_UNIV1";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1059
val Compl_Times_UNIV2 = thm "Compl_Times_UNIV2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1060
val PairE = thm "PairE";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1061
val Pair_Rep_inject = thm "Pair_Rep_inject";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1062
val Pair_def = thm "Pair_def";
27104
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 26975
diff changeset
  1063
val Pair_eq = @{thm "prod.inject"};
15404
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1064
val Pair_fst_snd_eq = thm "Pair_fst_snd_eq";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1065
val ProdI = thm "ProdI";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1066
val SetCompr_Sigma_eq = thm "SetCompr_Sigma_eq";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1067
val SigmaD1 = thm "SigmaD1";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1068
val SigmaD2 = thm "SigmaD2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1069
val SigmaE = thm "SigmaE";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1070
val SigmaE2 = thm "SigmaE2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1071
val SigmaI = thm "SigmaI";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1072
val Sigma_Diff_distrib1 = thm "Sigma_Diff_distrib1";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1073
val Sigma_Diff_distrib2 = thm "Sigma_Diff_distrib2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1074
val Sigma_Int_distrib1 = thm "Sigma_Int_distrib1";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1075
val Sigma_Int_distrib2 = thm "Sigma_Int_distrib2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1076
val Sigma_Un_distrib1 = thm "Sigma_Un_distrib1";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1077
val Sigma_Un_distrib2 = thm "Sigma_Un_distrib2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1078
val Sigma_Union = thm "Sigma_Union";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1079
val Sigma_def = thm "Sigma_def";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1080
val Sigma_empty1 = thm "Sigma_empty1";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1081
val Sigma_empty2 = thm "Sigma_empty2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1082
val Sigma_mono = thm "Sigma_mono";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1083
val The_split = thm "The_split";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1084
val The_split_eq = thm "The_split_eq";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1085
val The_split_eq = thm "The_split_eq";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1086
val Times_Diff_distrib1 = thm "Times_Diff_distrib1";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1087
val Times_Int_distrib1 = thm "Times_Int_distrib1";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1088
val Times_Un_distrib1 = thm "Times_Un_distrib1";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1089
val Times_eq_cancel2 = thm "Times_eq_cancel2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1090
val Times_subset_cancel2 = thm "Times_subset_cancel2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1091
val UNIV_Times_UNIV = thm "UNIV_Times_UNIV";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1092
val UN_Times_distrib = thm "UN_Times_distrib";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1093
val Unity_def = thm "Unity_def";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1094
val cond_split_eta = thm "cond_split_eta";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1095
val fst_conv = thm "fst_conv";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1096
val fst_def = thm "fst_def";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1097
val fst_eqD = thm "fst_eqD";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1098
val inj_on_Abs_Prod = thm "inj_on_Abs_Prod";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1099
val mem_Sigma_iff = thm "mem_Sigma_iff";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1100
val mem_splitE = thm "mem_splitE";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1101
val mem_splitI = thm "mem_splitI";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1102
val mem_splitI2 = thm "mem_splitI2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1103
val prod_eqI = thm "prod_eqI";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1104
val prod_fun = thm "prod_fun";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1105
val prod_fun_compose = thm "prod_fun_compose";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1106
val prod_fun_def = thm "prod_fun_def";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1107
val prod_fun_ident = thm "prod_fun_ident";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1108
val prod_fun_imageE = thm "prod_fun_imageE";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1109
val prod_fun_imageI = thm "prod_fun_imageI";
27104
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 26975
diff changeset
  1110
val prod_induct = thm "prod.induct";
15404
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1111
val snd_conv = thm "snd_conv";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1112
val snd_def = thm "snd_def";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1113
val snd_eqD = thm "snd_eqD";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1114
val split = thm "split";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1115
val splitD = thm "splitD";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1116
val splitD' = thm "splitD'";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1117
val splitE = thm "splitE";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1118
val splitE' = thm "splitE'";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1119
val splitE2 = thm "splitE2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1120
val splitI = thm "splitI";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1121
val splitI2 = thm "splitI2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1122
val splitI2' = thm "splitI2'";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1123
val split_beta = thm "split_beta";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1124
val split_conv = thm "split_conv";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1125
val split_def = thm "split_def";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1126
val split_eta = thm "split_eta";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1127
val split_eta_SetCompr = thm "split_eta_SetCompr";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1128
val split_eta_SetCompr2 = thm "split_eta_SetCompr2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1129
val split_paired_All = thm "split_paired_All";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1130
val split_paired_Ball_Sigma = thm "split_paired_Ball_Sigma";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1131
val split_paired_Bex_Sigma = thm "split_paired_Bex_Sigma";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1132
val split_paired_Ex = thm "split_paired_Ex";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1133
val split_paired_The = thm "split_paired_The";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1134
val split_paired_all = thm "split_paired_all";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1135
val split_part = thm "split_part";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1136
val split_split = thm "split_split";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1137
val split_split_asm = thm "split_split_asm";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1138
val split_tupled_all = thms "split_tupled_all";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1139
val split_weak_cong = thm "split_weak_cong";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1140
val surj_pair = thm "surj_pair";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1141
val surjective_pairing = thm "surjective_pairing";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1142
val unit_abs_eta_conv = thm "unit_abs_eta_conv";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1143
val unit_all_eq1 = thm "unit_all_eq1";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1144
val unit_all_eq2 = thm "unit_all_eq2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1145
val unit_eq = thm "unit_eq";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1146
*}
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15394
diff changeset
  1147
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
  1148
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
  1149
subsection {* Further inductive packages *}
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
  1150
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
  1151
use "Tools/inductive_realizer.ML"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
  1152
setup InductiveRealizer.setup
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
  1153
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31667
diff changeset
  1154
use "Tools/inductive_set.ML"
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31667
diff changeset
  1155
setup Inductive_Set.setup
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
  1156
31775
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 31723
diff changeset
  1157
use "Tools/Datatype/datatype_realizer.ML"
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
  1158
setup DatatypeRealizer.setup
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
  1159
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
  1160
end