src/ZF/pair.thy
author wenzelm
Thu, 30 Jan 2025 22:29:45 +0100
changeset 82023 9601f5582f33
parent 78099 4d9349989d94
permissions -rw-r--r--
more robust wrt. Par_List.map in Browser_Info.build(), see also 2fff9ce6b460 and 787a203a20b6;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41777
1f7cbe39d425 more precise headers;
wenzelm
parents: 28952
diff changeset
     1
(*  Title:      ZF/pair.thy
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
     3
    Copyright   1992  University of Cambridge
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
     4
*)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
     5
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59647
diff changeset
     6
section\<open>Ordered Pairs\<close>
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13240
diff changeset
     7
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14864
diff changeset
     8
theory pair imports upair
42455
6702c984bf5a modernized Quantifier1 simproc setup;
wenzelm
parents: 41777
diff changeset
     9
begin
6702c984bf5a modernized Quantifier1 simproc setup;
wenzelm
parents: 41777
diff changeset
    10
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
    11
ML_file \<open>simpdata.ML\<close>
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 46953
diff changeset
    12
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59647
diff changeset
    13
setup \<open>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48891
diff changeset
    14
  map_theory_simpset
60822
4f58f3662e7d more explicit context;
wenzelm
parents: 60770
diff changeset
    15
    (Simplifier.set_mksimps (fn ctxt => map mk_eq o ZF_atomize o Variable.gen_all ctxt)
45625
750c5a47400b modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 45620
diff changeset
    16
      #> Simplifier.add_cong @{thm if_weak_cong})
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59647
diff changeset
    17
\<close>
42794
07155da3b2f4 make ZF_cs snapshot after final setup;
wenzelm
parents: 42459
diff changeset
    18
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 60822
diff changeset
    19
ML \<open>val ZF_ss = simpset_of \<^context>\<close>
42794
07155da3b2f4 make ZF_cs snapshot after final setup;
wenzelm
parents: 42459
diff changeset
    20
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
    21
simproc_setup defined_Bex ("\<exists>x\<in>A. P(x) \<and> Q(x)") = \<open>
78099
4d9349989d94 more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
wenzelm
parents: 76216
diff changeset
    22
  K (Quantifier1.rearrange_Bex (fn ctxt => unfold_tac ctxt @{thms Bex_def}))
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59647
diff changeset
    23
\<close>
42455
6702c984bf5a modernized Quantifier1 simproc setup;
wenzelm
parents: 41777
diff changeset
    24
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59647
diff changeset
    25
simproc_setup defined_Ball ("\<forall>x\<in>A. P(x) \<longrightarrow> Q(x)") = \<open>
78099
4d9349989d94 more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
wenzelm
parents: 76216
diff changeset
    26
  K (Quantifier1.rearrange_Ball (fn ctxt => unfold_tac ctxt @{thms Ball_def}))
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59647
diff changeset
    27
\<close>
42455
6702c984bf5a modernized Quantifier1 simproc setup;
wenzelm
parents: 41777
diff changeset
    28
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    29
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    30
(** Lemmas for showing that \<langle>a,b\<rangle> uniquely determines a and b **)
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    31
46821
ff6b0c1087f2 Using mathematical notation for <-> and cardinal arithmetic
paulson
parents: 46820
diff changeset
    32
lemma singleton_eq_iff [iff]: "{a} = {b} \<longleftrightarrow> a=b"
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    33
by (rule extension [THEN iff_trans], blast)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    34
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
    35
lemma doubleton_eq_iff: "{a,b} = {c,d} \<longleftrightarrow> (a=c \<and> b=d) | (a=d \<and> b=c)"
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    36
by (rule extension [THEN iff_trans], blast)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    37
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    38
lemma Pair_iff [simp]: "\<langle>a,b\<rangle> = \<langle>c,d\<rangle> \<longleftrightarrow> a=c \<and> b=d"
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    39
by (simp add: Pair_def doubleton_eq_iff, blast)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    40
45602
2a858377c3d2 eliminated obsolete "standard";
wenzelm
parents: 42795
diff changeset
    41
lemmas Pair_inject = Pair_iff [THEN iffD1, THEN conjE, elim!]
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    42
45602
2a858377c3d2 eliminated obsolete "standard";
wenzelm
parents: 42795
diff changeset
    43
lemmas Pair_inject1 = Pair_iff [THEN iffD1, THEN conjunct1]
2a858377c3d2 eliminated obsolete "standard";
wenzelm
parents: 42795
diff changeset
    44
lemmas Pair_inject2 = Pair_iff [THEN iffD1, THEN conjunct2]
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    45
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    46
lemma Pair_not_0: "\<langle>a,b\<rangle> \<noteq> 0"
76216
9fc34f76b4e8 getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents: 76215
diff changeset
    47
  unfolding Pair_def
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    48
apply (blast elim: equalityE)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    49
done
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    50
45602
2a858377c3d2 eliminated obsolete "standard";
wenzelm
parents: 42795
diff changeset
    51
lemmas Pair_neq_0 = Pair_not_0 [THEN notE, elim!]
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    52
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    53
declare sym [THEN Pair_neq_0, elim!]
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    54
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    55
lemma Pair_neq_fst: "\<langle>a,b\<rangle>=a \<Longrightarrow> P"
46841
49b91b716cbe Structured and calculation-based proofs (with new trans rules!)
paulson
parents: 46821
diff changeset
    56
proof (unfold Pair_def)
49b91b716cbe Structured and calculation-based proofs (with new trans rules!)
paulson
parents: 46821
diff changeset
    57
  assume eq: "{{a, a}, {a, b}} = a"
49b91b716cbe Structured and calculation-based proofs (with new trans rules!)
paulson
parents: 46821
diff changeset
    58
  have  "{a, a} \<in> {{a, a}, {a, b}}" by (rule consI1)
49b91b716cbe Structured and calculation-based proofs (with new trans rules!)
paulson
parents: 46821
diff changeset
    59
  hence "{a, a} \<in> a" by (simp add: eq)
49b91b716cbe Structured and calculation-based proofs (with new trans rules!)
paulson
parents: 46821
diff changeset
    60
  moreover have "a \<in> {a, a}" by (rule consI1)
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46841
diff changeset
    61
  ultimately show "P" by (rule mem_asym)
46841
49b91b716cbe Structured and calculation-based proofs (with new trans rules!)
paulson
parents: 46821
diff changeset
    62
qed
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    63
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    64
lemma Pair_neq_snd: "\<langle>a,b\<rangle>=b \<Longrightarrow> P"
46841
49b91b716cbe Structured and calculation-based proofs (with new trans rules!)
paulson
parents: 46821
diff changeset
    65
proof (unfold Pair_def)
49b91b716cbe Structured and calculation-based proofs (with new trans rules!)
paulson
parents: 46821
diff changeset
    66
  assume eq: "{{a, a}, {a, b}} = b"
49b91b716cbe Structured and calculation-based proofs (with new trans rules!)
paulson
parents: 46821
diff changeset
    67
  have  "{a, b} \<in> {{a, a}, {a, b}}" by blast
49b91b716cbe Structured and calculation-based proofs (with new trans rules!)
paulson
parents: 46821
diff changeset
    68
  hence "{a, b} \<in> b" by (simp add: eq)
49b91b716cbe Structured and calculation-based proofs (with new trans rules!)
paulson
parents: 46821
diff changeset
    69
  moreover have "b \<in> {a, b}" by blast
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46841
diff changeset
    70
  ultimately show "P" by (rule mem_asym)
46841
49b91b716cbe Structured and calculation-based proofs (with new trans rules!)
paulson
parents: 46821
diff changeset
    71
qed
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    72
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    73
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59647
diff changeset
    74
subsection\<open>Sigma: Disjoint Union of a Family of Sets\<close>
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13240
diff changeset
    75
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59647
diff changeset
    76
text\<open>Generalizes Cartesian product\<close>
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    77
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    78
lemma Sigma_iff [simp]: "\<langle>a,b\<rangle>: Sigma(A,B) \<longleftrightarrow> a \<in> A \<and> b \<in> B(a)"
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    79
by (simp add: Sigma_def)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    80
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    81
lemma SigmaI [TC,intro!]: "\<lbrakk>a \<in> A;  b \<in> B(a)\<rbrakk> \<Longrightarrow> \<langle>a,b\<rangle> \<in> Sigma(A,B)"
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    82
by simp
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    83
45602
2a858377c3d2 eliminated obsolete "standard";
wenzelm
parents: 42795
diff changeset
    84
lemmas SigmaD1 = Sigma_iff [THEN iffD1, THEN conjunct1]
2a858377c3d2 eliminated obsolete "standard";
wenzelm
parents: 42795
diff changeset
    85
lemmas SigmaD2 = Sigma_iff [THEN iffD1, THEN conjunct2]
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    86
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    87
(*The general elimination rule*)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    88
lemma SigmaE [elim!]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71886
diff changeset
    89
    "\<lbrakk>c \<in> Sigma(A,B);
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    90
        \<And>x y.\<lbrakk>x \<in> A;  y \<in> B(x);  c=\<langle>x,y\<rangle>\<rbrakk> \<Longrightarrow> P
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71886
diff changeset
    91
\<rbrakk> \<Longrightarrow> P"
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46841
diff changeset
    92
by (unfold Sigma_def, blast)
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    93
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    94
lemma SigmaE2 [elim!]:
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    95
    "\<lbrakk>\<langle>a,b\<rangle> \<in> Sigma(A,B);
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71886
diff changeset
    96
        \<lbrakk>a \<in> A;  b \<in> B(a)\<rbrakk> \<Longrightarrow> P
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71886
diff changeset
    97
\<rbrakk> \<Longrightarrow> P"
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46841
diff changeset
    98
by (unfold Sigma_def, blast)
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
    99
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   100
lemma Sigma_cong:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71886
diff changeset
   101
    "\<lbrakk>A=A';  \<And>x. x \<in> A' \<Longrightarrow> B(x)=B'(x)\<rbrakk> \<Longrightarrow>
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   102
     Sigma(A,B) = Sigma(A',B')"
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   103
by (simp add: Sigma_def)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   104
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   105
(*Sigma_cong, Pi_cong NOT given to Addcongs: they cause
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   106
  flex-flex pairs and the "Check your prover" error.  Most
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   107
  Sigmas and Pis are abbreviated as * or -> *)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   108
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   109
lemma Sigma_empty1 [simp]: "Sigma(0,B) = 0"
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   110
by blast
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   111
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   112
lemma Sigma_empty2 [simp]: "A*0 = 0"
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   113
by blast
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   114
46821
ff6b0c1087f2 Using mathematical notation for <-> and cardinal arithmetic
paulson
parents: 46820
diff changeset
   115
lemma Sigma_empty_iff: "A*B=0 \<longleftrightarrow> A=0 | B=0"
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   116
by blast
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   117
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   118
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 60822
diff changeset
   119
subsection\<open>Projections \<^term>\<open>fst\<close> and \<^term>\<open>snd\<close>\<close>
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   120
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   121
lemma fst_conv [simp]: "fst(\<langle>a,b\<rangle>) = a"
13544
895994073bdf various new lemmas for Constructible
paulson
parents: 13357
diff changeset
   122
by (simp add: fst_def)
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   123
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   124
lemma snd_conv [simp]: "snd(\<langle>a,b\<rangle>) = b"
13544
895994073bdf various new lemmas for Constructible
paulson
parents: 13357
diff changeset
   125
by (simp add: snd_def)
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   126
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71886
diff changeset
   127
lemma fst_type [TC]: "p \<in> Sigma(A,B) \<Longrightarrow> fst(p) \<in> A"
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   128
by auto
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   129
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71886
diff changeset
   130
lemma snd_type [TC]: "p \<in> Sigma(A,B) \<Longrightarrow> snd(p) \<in> B(fst(p))"
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   131
by auto
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   132
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71886
diff changeset
   133
lemma Pair_fst_snd_eq: "a \<in> Sigma(A,B) \<Longrightarrow> <fst(a),snd(a)> = a"
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   134
by auto
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   135
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   136
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 60822
diff changeset
   137
subsection\<open>The Eliminator, \<^term>\<open>split\<close>\<close>
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   138
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   139
(*A META-equality, so that it applies to higher types as well...*)
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   140
lemma split [simp]: "split(\<lambda>x y. c(x,y), \<langle>a,b\<rangle>) \<equiv> c(a,b)"
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   141
by (simp add: split_def)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   142
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   143
lemma split_type [TC]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71886
diff changeset
   144
    "\<lbrakk>p \<in> Sigma(A,B);
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   145
         \<And>x y.\<lbrakk>x \<in> A; y \<in> B(x)\<rbrakk> \<Longrightarrow> c(x,y):C(\<langle>x,y\<rangle>)
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   146
\<rbrakk> \<Longrightarrow> split(\<lambda>x y. c(x,y), p) \<in> C(p)"
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46841
diff changeset
   147
by (erule SigmaE, auto)
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   148
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46841
diff changeset
   149
lemma expand_split:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71886
diff changeset
   150
  "u \<in> A*B \<Longrightarrow>
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   151
        R(split(c,u)) \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>B. u = \<langle>x,y\<rangle> \<longrightarrow> R(c(x,y)))"
46841
49b91b716cbe Structured and calculation-based proofs (with new trans rules!)
paulson
parents: 46821
diff changeset
   152
by (auto simp add: split_def)
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   153
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   154
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 60822
diff changeset
   155
subsection\<open>A version of \<^term>\<open>split\<close> for Formulae: Result Type \<^typ>\<open>o\<close>\<close>
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   156
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   157
lemma splitI: "R(a,b) \<Longrightarrow> split(R, \<langle>a,b\<rangle>)"
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   158
by (simp add: split_def)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   159
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   160
lemma splitE:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71886
diff changeset
   161
    "\<lbrakk>split(R,z);  z \<in> Sigma(A,B);
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   162
        \<And>x y. \<lbrakk>z = \<langle>x,y\<rangle>;  R(x,y)\<rbrakk> \<Longrightarrow> P
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71886
diff changeset
   163
\<rbrakk> \<Longrightarrow> P"
46841
49b91b716cbe Structured and calculation-based proofs (with new trans rules!)
paulson
parents: 46821
diff changeset
   164
by (auto simp add: split_def)
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   165
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   166
lemma splitD: "split(R,\<langle>a,b\<rangle>) \<Longrightarrow> R(a,b)"
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   167
by (simp add: split_def)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 11694
diff changeset
   168
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59647
diff changeset
   169
text \<open>
14864
419b45cdb400 new rules for simplifying quantifiers with Sigma
paulson
parents: 13544
diff changeset
   170
  \bigskip Complex rules for Sigma.
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59647
diff changeset
   171
\<close>
14864
419b45cdb400 new rules for simplifying quantifiers with Sigma
paulson
parents: 13544
diff changeset
   172
419b45cdb400 new rules for simplifying quantifiers with Sigma
paulson
parents: 13544
diff changeset
   173
lemma split_paired_Bex_Sigma [simp]:
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   174
     "(\<exists>z \<in> Sigma(A,B). P(z)) \<longleftrightarrow> (\<exists>x \<in> A. \<exists>y \<in> B(x). P(\<langle>x,y\<rangle>))"
14864
419b45cdb400 new rules for simplifying quantifiers with Sigma
paulson
parents: 13544
diff changeset
   175
by blast
419b45cdb400 new rules for simplifying quantifiers with Sigma
paulson
parents: 13544
diff changeset
   176
419b45cdb400 new rules for simplifying quantifiers with Sigma
paulson
parents: 13544
diff changeset
   177
lemma split_paired_Ball_Sigma [simp]:
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   178
     "(\<forall>z \<in> Sigma(A,B). P(z)) \<longleftrightarrow> (\<forall>x \<in> A. \<forall>y \<in> B(x). P(\<langle>x,y\<rangle>))"
14864
419b45cdb400 new rules for simplifying quantifiers with Sigma
paulson
parents: 13544
diff changeset
   179
by blast
419b45cdb400 new rules for simplifying quantifiers with Sigma
paulson
parents: 13544
diff changeset
   180
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 2469
diff changeset
   181
end
124
858ab9a9b047 made pseudo theories for all ML files;
clasohm
parents:
diff changeset
   182
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 124
diff changeset
   183