src/ZF/func.thy
author wenzelm
Sun, 23 Apr 2023 21:31:00 +0200
changeset 77912 430e6c477ba4
parent 76217 8655344f1cf6
child 80754 701912f5645a
permissions -rw-r--r--
more operations for lexicographic ordering;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
     1
(*  Title:      ZF/func.thy
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
     3
    Copyright   1991  University of Cambridge
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
     4
*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
     5
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
     6
section\<open>Functions, Function Spaces, Lambda-Abstraction\<close>
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
     7
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14883
diff changeset
     8
theory func imports equalities Sum begin
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
     9
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
    10
subsection\<open>The Pi Operator: Dependent Function Space\<close>
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
    11
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
    12
lemma subset_Sigma_imp_relation: "r \<subseteq> Sigma(A,B) \<Longrightarrow> relation(r)"
13248
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
    13
by (simp add: relation_def, blast)
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
    14
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 13219
diff changeset
    15
lemma relation_converse_converse [simp]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
    16
     "relation(r) \<Longrightarrow> converse(converse(r)) = r"
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
    17
by (simp add: relation_def, blast)
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 13219
diff changeset
    18
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 13219
diff changeset
    19
lemma relation_restrict [simp]:  "relation(restrict(r,A))"
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
    20
by (simp add: restrict_def relation_def, blast)
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 13219
diff changeset
    21
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    22
lemma Pi_iff:
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
    23
    "f \<in> Pi(A,B) \<longleftrightarrow> function(f) \<and> f<=Sigma(A,B) \<and> A<=domain(f)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    24
by (unfold Pi_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    25
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    26
(*For upward compatibility with the former definition*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    27
lemma Pi_iff_old:
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    28
    "f \<in> Pi(A,B) \<longleftrightarrow> f<=Sigma(A,B) \<and> (\<forall>x\<in>A. \<exists>!y. \<langle>x,y\<rangle>: f)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    29
by (unfold Pi_def function_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    30
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
    31
lemma fun_is_function: "f \<in> Pi(A,B) \<Longrightarrow> function(f)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    32
by (simp only: Pi_iff)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    33
13219
7e44aa8a276e new lemma
paulson
parents: 13179
diff changeset
    34
lemma function_imp_Pi:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
    35
     "\<lbrakk>function(f); relation(f)\<rbrakk> \<Longrightarrow> f \<in> domain(f) -> range(f)"
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
    36
by (simp add: Pi_iff relation_def, blast)
13219
7e44aa8a276e new lemma
paulson
parents: 13179
diff changeset
    37
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
    38
lemma functionI:
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    39
     "\<lbrakk>\<And>x y y'. \<lbrakk>\<langle>x,y\<rangle>:r; <x,y'>:r\<rbrakk> \<Longrightarrow> y=y'\<rbrakk> \<Longrightarrow> function(r)"
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
    40
by (simp add: function_def, blast)
13172
03a5afa7b888 tidying up
paulson
parents: 13168
diff changeset
    41
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    42
(*Functions are relations*)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
    43
lemma fun_is_rel: "f \<in> Pi(A,B) \<Longrightarrow> f \<subseteq> Sigma(A,B)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    44
by (unfold Pi_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    45
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    46
lemma Pi_cong:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
    47
    "\<lbrakk>A=A';  \<And>x. x \<in> A' \<Longrightarrow> B(x)=B'(x)\<rbrakk> \<Longrightarrow> Pi(A,B) = Pi(A',B')"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    48
by (simp add: Pi_def cong add: Sigma_cong)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    49
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    50
(*Sigma_cong, Pi_cong NOT given to Addcongs: they cause
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    51
  flex-flex pairs and the "Check your prover" error.  Most
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    52
  Sigmas and Pis are abbreviated as * or -> *)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    53
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    54
(*Weakening one function type to another; see also Pi_type*)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
    55
lemma fun_weaken_type: "\<lbrakk>f \<in> A->B;  B<=D\<rbrakk> \<Longrightarrow> f \<in> A->D"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    56
by (unfold Pi_def, best)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    57
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
    58
subsection\<open>Function Application\<close>
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    59
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    60
lemma apply_equality2: "\<lbrakk>\<langle>a,b\<rangle>: f;  \<langle>a,c\<rangle>: f;  f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> b=c"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    61
by (unfold Pi_def function_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    62
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    63
lemma function_apply_equality: "\<lbrakk>\<langle>a,b\<rangle>: f;  function(f)\<rbrakk> \<Longrightarrow> f`a = b"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    64
by (unfold apply_def function_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    65
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    66
lemma apply_equality: "\<lbrakk>\<langle>a,b\<rangle>: f;  f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> f`a = b"
76216
9fc34f76b4e8 getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents: 76215
diff changeset
    67
  unfolding Pi_def
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    68
apply (blast intro: function_apply_equality)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    69
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    70
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    71
(*Applying a function outside its domain yields 0*)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
    72
lemma apply_0: "a \<notin> domain(f) \<Longrightarrow> f`a = 0"
13176
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
    73
by (unfold apply_def, blast)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    74
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
    75
lemma Pi_memberD: "\<lbrakk>f \<in> Pi(A,B);  c \<in> f\<rbrakk> \<Longrightarrow> \<exists>x\<in>A.  c = <x,f`x>"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    76
apply (frule fun_is_rel)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    77
apply (blast dest: apply_equality)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    78
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    79
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
    80
lemma function_apply_Pair: "\<lbrakk>function(f);  a \<in> domain(f)\<rbrakk> \<Longrightarrow> <a,f`a>: f"
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
    81
apply (simp add: function_def, clarify)
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
    82
apply (subgoal_tac "f`a = y", blast)
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
    83
apply (simp add: apply_def, blast)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    84
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    85
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
    86
lemma apply_Pair: "\<lbrakk>f \<in> Pi(A,B);  a \<in> A\<rbrakk> \<Longrightarrow> <a,f`a>: f"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    87
apply (simp add: Pi_iff)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    88
apply (blast intro: function_apply_Pair)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    89
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    90
27150
a42aef558ce3 tuned comments;
wenzelm
parents: 24893
diff changeset
    91
(*Conclusion is flexible -- use rule_tac or else apply_funtype below!*)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
    92
lemma apply_type [TC]: "\<lbrakk>f \<in> Pi(A,B);  a \<in> A\<rbrakk> \<Longrightarrow> f`a \<in> B(a)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    93
by (blast intro: apply_Pair dest: fun_is_rel)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    94
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    95
(*This version is acceptable to the simplifier*)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
    96
lemma apply_funtype: "\<lbrakk>f \<in> A->B;  a \<in> A\<rbrakk> \<Longrightarrow> f`a \<in> B"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    97
by (blast dest: apply_type)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    98
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    99
lemma apply_iff: "f \<in> Pi(A,B) \<Longrightarrow> \<langle>a,b\<rangle>: f \<longleftrightarrow> a \<in> A \<and> f`a = b"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   100
apply (frule fun_is_rel)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   101
apply (blast intro!: apply_Pair apply_equality)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   102
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   103
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   104
(*Refining one Pi type to another*)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   105
lemma Pi_type: "\<lbrakk>f \<in> Pi(A,C);  \<And>x. x \<in> A \<Longrightarrow> f`x \<in> B(x)\<rbrakk> \<Longrightarrow> f \<in> Pi(A,B)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   106
apply (simp only: Pi_iff)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   107
apply (blast dest: function_apply_equality)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   108
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   109
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   110
(*Such functions arise in non-standard datatypes, ZF/ex/Ntree for instance*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   111
lemma Pi_Collect_iff:
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   112
     "(f \<in> Pi(A, \<lambda>x. {y \<in> B(x). P(x,y)}))
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   113
      \<longleftrightarrow>  f \<in> Pi(A,B) \<and> (\<forall>x\<in>A. P(x, f`x))"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   114
by (blast intro: Pi_type dest: apply_type)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   115
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   116
lemma Pi_weaken_type:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   117
        "\<lbrakk>f \<in> Pi(A,B);  \<And>x. x \<in> A \<Longrightarrow> B(x)<=C(x)\<rbrakk> \<Longrightarrow> f \<in> Pi(A,C)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   118
by (blast intro: Pi_type dest: apply_type)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   119
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   120
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   121
(** Elimination of membership in a function **)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   122
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   123
lemma domain_type: "\<lbrakk>\<langle>a,b\<rangle> \<in> f;  f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> a \<in> A"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   124
by (blast dest: fun_is_rel)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   125
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   126
lemma range_type: "\<lbrakk>\<langle>a,b\<rangle> \<in> f;  f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> b \<in> B(a)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   127
by (blast dest: fun_is_rel)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   128
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   129
lemma Pair_mem_PiD: "\<lbrakk>\<langle>a,b\<rangle>: f;  f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> a \<in> A \<and> b \<in> B(a) \<and> f`a = b"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   130
by (blast intro: domain_type range_type apply_equality)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   131
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   132
subsection\<open>Lambda Abstraction\<close>
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   133
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   134
lemma lamI: "a \<in> A \<Longrightarrow> <a,b(a)> \<in> (\<lambda>x\<in>A. b(x))"
76216
9fc34f76b4e8 getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents: 76215
diff changeset
   135
  unfolding lam_def
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   136
apply (erule RepFunI)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   137
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   138
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   139
lemma lamE:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   140
    "\<lbrakk>p: (\<lambda>x\<in>A. b(x));  \<And>x.\<lbrakk>x \<in> A; p=<x,b(x)>\<rbrakk> \<Longrightarrow> P
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   141
\<rbrakk> \<Longrightarrow>  P"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   142
by (simp add: lam_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   143
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   144
lemma lamD: "\<lbrakk>\<langle>a,c\<rangle>: (\<lambda>x\<in>A. b(x))\<rbrakk> \<Longrightarrow> c = b(a)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   145
by (simp add: lam_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   146
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   147
lemma lam_type [TC]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   148
    "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> b(x): B(x)\<rbrakk> \<Longrightarrow> (\<lambda>x\<in>A. b(x)) \<in> Pi(A,B)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   149
by (simp add: lam_def Pi_def function_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   150
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   151
lemma lam_funtype: "(\<lambda>x\<in>A. b(x)) \<in> A -> {b(x). x \<in> A}"
13176
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
   152
by (blast intro: lam_type)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   153
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 41229
diff changeset
   154
lemma function_lam: "function (\<lambda>x\<in>A. b(x))"
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   155
by (simp add: function_def lam_def)
13172
03a5afa7b888 tidying up
paulson
parents: 13168
diff changeset
   156
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   157
lemma relation_lam: "relation (\<lambda>x\<in>A. b(x))"
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   158
by (simp add: relation_def lam_def)
13172
03a5afa7b888 tidying up
paulson
parents: 13168
diff changeset
   159
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 41229
diff changeset
   160
lemma beta_if [simp]: "(\<lambda>x\<in>A. b(x)) ` a = (if a \<in> A then b(a) else 0)"
13176
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
   161
by (simp add: apply_def lam_def, blast)
13175
81082cfa5618 new definition of "apply" and new simprule "beta_if"
paulson
parents: 13174
diff changeset
   162
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   163
lemma beta: "a \<in> A \<Longrightarrow> (\<lambda>x\<in>A. b(x)) ` a = b(a)"
13176
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
   164
by (simp add: apply_def lam_def, blast)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   165
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 41229
diff changeset
   166
lemma lam_empty [simp]: "(\<lambda>x\<in>0. b(x)) = 0"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   167
by (simp add: lam_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   168
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   169
lemma domain_lam [simp]: "domain(Lambda(A,b)) = A"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   170
by (simp add: lam_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   171
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   172
(*congruence rule for lambda abstraction*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   173
lemma lam_cong [cong]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   174
    "\<lbrakk>A=A';  \<And>x. x \<in> A' \<Longrightarrow> b(x)=b'(x)\<rbrakk> \<Longrightarrow> Lambda(A,b) = Lambda(A',b')"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   175
by (simp only: lam_def cong add: RepFun_cong)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   176
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   177
lemma lam_theI:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   178
    "(\<And>x. x \<in> A \<Longrightarrow> \<exists>!y. Q(x,y)) \<Longrightarrow> \<exists>f. \<forall>x\<in>A. Q(x, f`x)"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 41229
diff changeset
   179
apply (rule_tac x = "\<lambda>x\<in>A. THE y. Q (x,y)" in exI)
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   180
apply simp
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   181
apply (blast intro: theI)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   182
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   183
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   184
lemma lam_eqE: "\<lbrakk>(\<lambda>x\<in>A. f(x)) = (\<lambda>x\<in>A. g(x));  a \<in> A\<rbrakk> \<Longrightarrow> f(a)=g(a)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   185
by (fast intro!: lamI elim: equalityE lamE)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   186
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   187
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   188
(*Empty function spaces*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   189
lemma Pi_empty1 [simp]: "Pi(0,A) = {0}"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   190
by (unfold Pi_def function_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   191
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   192
(*The singleton function*)
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   193
lemma singleton_fun [simp]: "{\<langle>a,b\<rangle>} \<in> {a} -> {b}"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   194
by (unfold Pi_def function_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   195
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   196
lemma Pi_empty2 [simp]: "(A->0) = (if A=0 then {0} else 0)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   197
by (unfold Pi_def function_def, force)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   198
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   199
lemma  fun_space_empty_iff [iff]: "(A->X)=0 \<longleftrightarrow> X=0 \<and> (A \<noteq> 0)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   200
apply auto
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   201
apply (fast intro!: equals0I intro: lam_type)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   202
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   203
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   204
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   205
subsection\<open>Extensionality\<close>
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   206
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   207
(*Semi-extensionality!*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   208
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   209
lemma fun_subset:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   210
    "\<lbrakk>f \<in> Pi(A,B);  g \<in> Pi(C,D);  A<=C;
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   211
        \<And>x. x \<in> A \<Longrightarrow> f`x = g`x\<rbrakk> \<Longrightarrow> f<=g"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   212
by (force dest: Pi_memberD intro: apply_Pair)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   213
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   214
lemma fun_extension:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   215
    "\<lbrakk>f \<in> Pi(A,B);  g \<in> Pi(A,D);
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   216
        \<And>x. x \<in> A \<Longrightarrow> f`x = g`x\<rbrakk> \<Longrightarrow> f=g"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   217
by (blast del: subsetI intro: subset_refl sym fun_subset)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   218
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   219
lemma eta [simp]: "f \<in> Pi(A,B) \<Longrightarrow> (\<lambda>x\<in>A. f`x) = f"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   220
apply (rule fun_extension)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   221
apply (auto simp add: lam_type apply_type beta)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   222
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   223
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   224
lemma fun_extension_iff:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   225
     "\<lbrakk>f \<in> Pi(A,B); g \<in> Pi(A,C)\<rbrakk> \<Longrightarrow> (\<forall>a\<in>A. f`a = g`a) \<longleftrightarrow> f=g"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   226
by (blast intro: fun_extension)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   227
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   228
(*thm by Mark Staples, proof by lcp*)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   229
lemma fun_subset_eq: "\<lbrakk>f \<in> Pi(A,B); g \<in> Pi(A,C)\<rbrakk> \<Longrightarrow> f \<subseteq> g \<longleftrightarrow> (f = g)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   230
by (blast dest: apply_Pair
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 27702
diff changeset
   231
          intro: fun_extension apply_equality [symmetric])
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   232
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   233
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   234
(*Every element of Pi(A,B) may be expressed as a lambda abstraction!*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   235
lemma Pi_lamE:
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   236
  assumes major: "f \<in> Pi(A,B)"
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   237
      and minor: "\<And>b. \<lbrakk>\<forall>x\<in>A. b(x):B(x);  f = (\<lambda>x\<in>A. b(x))\<rbrakk> \<Longrightarrow> P"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   238
  shows "P"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   239
apply (rule minor)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   240
apply (rule_tac [2] eta [symmetric])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   241
apply (blast intro: major apply_type)+
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   242
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   243
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   244
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   245
subsection\<open>Images of Functions\<close>
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   246
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   247
lemma image_lam: "C \<subseteq> A \<Longrightarrow> (\<lambda>x\<in>A. b(x)) `` C = {b(x). x \<in> C}"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   248
by (unfold lam_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   249
13179
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   250
lemma Repfun_function_if:
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   251
     "function(f)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   252
      \<Longrightarrow> {f`x. x \<in> C} = (if C \<subseteq> domain(f) then f``C else cons(0,f``C))"
13179
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   253
apply simp
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   254
apply (intro conjI impI)
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   255
 apply (blast dest: function_apply_equality intro: function_apply_Pair)
13179
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   256
apply (rule equalityI)
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   257
 apply (blast intro!: function_apply_Pair apply_0)
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   258
apply (blast dest: function_apply_equality intro: apply_0 [symmetric])
13179
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   259
done
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   260
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   261
(*For this lemma and the next, the right-hand side could equivalently
13615
449a70d88b38 Numerous cosmetic changes, prompted by the new simplifier
paulson
parents: 13357
diff changeset
   262
  be written \<Union>x\<in>C. {f`x} *)
13174
85d3c0981a16 more tidying
paulson
parents: 13172
diff changeset
   263
lemma image_function:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   264
     "\<lbrakk>function(f);  C \<subseteq> domain(f)\<rbrakk> \<Longrightarrow> f``C = {f`x. x \<in> C}"
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   265
by (simp add: Repfun_function_if)
13174
85d3c0981a16 more tidying
paulson
parents: 13172
diff changeset
   266
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   267
lemma image_fun: "\<lbrakk>f \<in> Pi(A,B);  C \<subseteq> A\<rbrakk> \<Longrightarrow> f``C = {f`x. x \<in> C}"
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   268
apply (simp add: Pi_iff)
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   269
apply (blast intro: image_function)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   270
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   271
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   272
lemma image_eq_UN:
14883
ca000a495448 Groups, Rings and supporting lemmas
paulson
parents: 14153
diff changeset
   273
  assumes f: "f \<in> Pi(A,B)" "C \<subseteq> A" shows "f``C = (\<Union>x\<in>C. {f ` x})"
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   274
by (auto simp add: image_fun [OF f])
14883
ca000a495448 Groups, Rings and supporting lemmas
paulson
parents: 14153
diff changeset
   275
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   276
lemma Pi_image_cons:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   277
     "\<lbrakk>f \<in> Pi(A,B);  x \<in> A\<rbrakk> \<Longrightarrow> f `` cons(x,y) = cons(f`x, f``y)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   278
by (blast dest: apply_equality apply_Pair)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   279
124
858ab9a9b047 made pseudo theories for all ML files;
clasohm
parents:
diff changeset
   280
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69587
diff changeset
   281
subsection\<open>Properties of \<^term>\<open>restrict(f,A)\<close>\<close>
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   282
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 41229
diff changeset
   283
lemma restrict_subset: "restrict(f,A) \<subseteq> f"
13179
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   284
by (unfold restrict_def, blast)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   285
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   286
lemma function_restrictI:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   287
    "function(f) \<Longrightarrow> function(restrict(f,A))"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   288
by (unfold restrict_def function_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   289
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   290
lemma restrict_type2: "\<lbrakk>f \<in> Pi(C,B);  A<=C\<rbrakk> \<Longrightarrow> restrict(f,A) \<in> Pi(A,B)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   291
by (simp add: Pi_iff function_def restrict_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   292
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 41229
diff changeset
   293
lemma restrict: "restrict(f,A) ` a = (if a \<in> A then f`a else 0)"
13176
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
   294
by (simp add: apply_def restrict_def, blast)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   295
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   296
lemma restrict_empty [simp]: "restrict(f,0) = 0"
13179
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   297
by (unfold restrict_def, simp)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   298
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   299
lemma restrict_iff: "z \<in> restrict(r,A) \<longleftrightarrow> z \<in> r \<and> (\<exists>x\<in>A. \<exists>y. z = \<langle>x, y\<rangle>)"
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   300
by (simp add: restrict_def)
13172
03a5afa7b888 tidying up
paulson
parents: 13168
diff changeset
   301
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   302
lemma restrict_restrict [simp]:
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 41229
diff changeset
   303
     "restrict(restrict(r,A),B) = restrict(r, A \<inter> B)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   304
by (unfold restrict_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   305
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 41229
diff changeset
   306
lemma domain_restrict [simp]: "domain(restrict(f,C)) = domain(f) \<inter> C"
76216
9fc34f76b4e8 getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents: 76215
diff changeset
   307
  unfolding restrict_def
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   308
apply (auto simp add: domain_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   309
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   310
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   311
lemma restrict_idem: "f \<subseteq> Sigma(A,B) \<Longrightarrow> restrict(f,A) = f"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   312
by (simp add: restrict_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   313
13248
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   314
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   315
(*converse probably holds too*)
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   316
lemma domain_restrict_idem:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   317
     "\<lbrakk>domain(r) \<subseteq> A; relation(r)\<rbrakk> \<Longrightarrow> restrict(r,A) = r"
13248
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   318
by (simp add: restrict_def relation_def, blast)
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   319
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 41229
diff changeset
   320
lemma domain_restrict_lam [simp]: "domain(restrict(Lambda(A,f),C)) = A \<inter> C"
76217
8655344f1cf6 More obsolete "unfold" calls
paulson <lp15@cam.ac.uk>
parents: 76216
diff changeset
   321
  unfolding restrict_def lam_def
13248
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   322
apply (rule equalityI)
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   323
apply (auto simp add: domain_iff)
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   324
done
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   325
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 41229
diff changeset
   326
lemma restrict_if [simp]: "restrict(f,A) ` a = (if a \<in> A then f`a else 0)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   327
by (simp add: restrict apply_0)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   328
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   329
lemma restrict_lam_eq:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   330
    "A<=C \<Longrightarrow> restrict(\<lambda>x\<in>C. b(x), A) = (\<lambda>x\<in>A. b(x))"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   331
by (unfold restrict_def lam_def, auto)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   332
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   333
lemma fun_cons_restrict_eq:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   334
     "f \<in> cons(a, b) -> B \<Longrightarrow> f = cons(<a, f ` a>, restrict(f, b))"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   335
apply (rule equalityI)
13248
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   336
 prefer 2 apply (blast intro: apply_Pair restrict_subset [THEN subsetD])
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   337
apply (auto dest!: Pi_memberD simp add: restrict_def lam_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   338
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   339
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   340
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   341
subsection\<open>Unions of Functions\<close>
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   342
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   343
(** The Union of a set of COMPATIBLE functions is a function **)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   344
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   345
lemma function_Union:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   346
    "\<lbrakk>\<forall>x\<in>S. function(x);
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   347
        \<forall>x\<in>S. \<forall>y\<in>S. x<=y | y<=x\<rbrakk>
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   348
     \<Longrightarrow> function(\<Union>(S))"
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   349
by (unfold function_def, blast)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   350
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   351
lemma fun_Union:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   352
    "\<lbrakk>\<forall>f\<in>S. \<exists>C D. f \<in> C->D;
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   353
             \<forall>f\<in>S. \<forall>y\<in>S. f<=y | y<=f\<rbrakk> \<Longrightarrow>
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 41229
diff changeset
   354
          \<Union>(S) \<in> domain(\<Union>(S)) -> range(\<Union>(S))"
76216
9fc34f76b4e8 getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents: 76215
diff changeset
   355
  unfolding Pi_def
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   356
apply (blast intro!: rel_Union function_Union)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   357
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   358
71085
950e1cfe0fe4 tuned proofs -- more stable proof terms without [rule_format];
wenzelm
parents: 69593
diff changeset
   359
lemma gen_relation_Union:
950e1cfe0fe4 tuned proofs -- more stable proof terms without [rule_format];
wenzelm
parents: 69593
diff changeset
   360
     "(\<And>f. f\<in>F \<Longrightarrow> relation(f)) \<Longrightarrow> relation(\<Union>(F))"
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   361
by (simp add: relation_def)
13174
85d3c0981a16 more tidying
paulson
parents: 13172
diff changeset
   362
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   363
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   364
(** The Union of 2 disjoint functions is a function **)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   365
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   366
lemmas Un_rls = Un_subset_iff SUM_Un_distrib1 prod_Un_distrib2
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   367
                subset_trans [OF _ Un_upper1]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   368
                subset_trans [OF _ Un_upper2]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   369
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   370
lemma fun_disjoint_Un:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   371
     "\<lbrakk>f \<in> A->B;  g \<in> C->D;  A \<inter> C = 0\<rbrakk>
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   372
      \<Longrightarrow> (f \<union> g) \<in> (A \<union> C) -> (B \<union> D)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   373
(*Prove the product and domain subgoals using distributive laws*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   374
apply (simp add: Pi_iff extension Un_rls)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   375
apply (unfold function_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   376
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   377
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   378
lemma fun_disjoint_apply1: "a \<notin> domain(g) \<Longrightarrow> (f \<union> g)`a = f`a"
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   379
by (simp add: apply_def, blast)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   380
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   381
lemma fun_disjoint_apply2: "c \<notin> domain(f) \<Longrightarrow> (f \<union> g)`c = g`c"
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   382
by (simp add: apply_def, blast)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   383
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   384
subsection\<open>Domain and Range of a Function or Relation\<close>
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   385
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   386
lemma domain_of_fun: "f \<in> Pi(A,B) \<Longrightarrow> domain(f)=A"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   387
by (unfold Pi_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   388
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   389
lemma apply_rangeI: "\<lbrakk>f \<in> Pi(A,B);  a \<in> A\<rbrakk> \<Longrightarrow> f`a \<in> range(f)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   390
by (erule apply_Pair [THEN rangeI], assumption)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   391
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   392
lemma range_of_fun: "f \<in> Pi(A,B) \<Longrightarrow> f \<in> A->range(f)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   393
by (blast intro: Pi_type apply_rangeI)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   394
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   395
subsection\<open>Extensions of Functions\<close>
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   396
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   397
lemma fun_extend:
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   398
     "\<lbrakk>f \<in> A->B;  c\<notin>A\<rbrakk> \<Longrightarrow> cons(\<langle>c,b\<rangle>,f) \<in> cons(c,A) -> cons(b,B)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   399
apply (frule singleton_fun [THEN fun_disjoint_Un], blast)
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   400
apply (simp add: cons_eq)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   401
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   402
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   403
lemma fun_extend3:
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   404
     "\<lbrakk>f \<in> A->B;  c\<notin>A;  b \<in> B\<rbrakk> \<Longrightarrow> cons(\<langle>c,b\<rangle>,f) \<in> cons(c,A) -> B"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   405
by (blast intro: fun_extend [THEN fun_weaken_type])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   406
13176
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
   407
lemma extend_apply:
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   408
     "c \<notin> domain(f) \<Longrightarrow> cons(\<langle>c,b\<rangle>,f)`a = (if a=c then b else f`a)"
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   409
by (auto simp add: apply_def)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   410
13176
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
   411
lemma fun_extend_apply [simp]:
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   412
     "\<lbrakk>f \<in> A->B;  c\<notin>A\<rbrakk> \<Longrightarrow> cons(\<langle>c,b\<rangle>,f)`a = (if a=c then b else f`a)"
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   413
apply (rule extend_apply)
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   414
apply (simp add: Pi_def, blast)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   415
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   416
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   417
lemmas singleton_apply = apply_equality [OF singletonI singleton_fun, simp]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   418
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   419
(*For Finite.ML.  Inclusion of right into left is easy*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   420
lemma cons_fun_eq:
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   421
     "c \<notin> A \<Longrightarrow> cons(c,A) -> B = (\<Union>f \<in> A->B. \<Union>b\<in>B. {cons(\<langle>c,b\<rangle>, f)})"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   422
apply (rule equalityI)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   423
apply (safe elim!: fun_extend3)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   424
(*Inclusion of left into right*)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 41229
diff changeset
   425
apply (subgoal_tac "restrict (x, A) \<in> A -> B")
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   426
 prefer 2 apply (blast intro: restrict_type2)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   427
apply (rule UN_I, assumption)
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   428
apply (rule apply_funtype [THEN UN_I])
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   429
  apply assumption
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   430
 apply (rule consI1)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   431
apply (simp (no_asm))
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   432
apply (rule fun_extension)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   433
  apply assumption
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   434
 apply (blast intro: fun_extend)
13176
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
   435
apply (erule consE, simp_all)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   436
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   437
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   438
lemma succ_fun_eq: "succ(n) -> B = (\<Union>f \<in> n->B. \<Union>b\<in>B. {cons(\<langle>n,b\<rangle>, f)})"
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13248
diff changeset
   439
by (simp add: succ_def mem_not_refl cons_fun_eq)
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13248
diff changeset
   440
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   441
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   442
subsection\<open>Function Updates\<close>
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   443
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
   444
definition
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   445
  update  :: "[i,i,i] \<Rightarrow> i"  where
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   446
   "update(f,a,b) \<equiv> \<lambda>x\<in>cons(a, domain(f)). if(x=a, b, f`x)"
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   447
41229
d797baa3d57c replaced command 'nonterminals' by slightly modernized version 'nonterminal';
wenzelm
parents: 32960
diff changeset
   448
nonterminal updbinds and updbind
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   449
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   450
syntax
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   451
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   452
  (* Let expressions *)
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   453
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   454
  "_updbind"    :: "[i, i] \<Rightarrow> updbind"               (\<open>(2_ :=/ _)\<close>)
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   455
  ""            :: "updbind \<Rightarrow> updbinds"             (\<open>_\<close>)
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   456
  "_updbinds"   :: "[updbind, updbinds] \<Rightarrow> updbinds" (\<open>_,/ _\<close>)
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   457
  "_Update"     :: "[i, updbinds] \<Rightarrow> i"              (\<open>_/'((_)')\<close> [900,0] 900)
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   458
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   459
translations
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   460
  "_Update (f, _updbinds(b,bs))"  == "_Update (_Update(f,b), bs)"
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
   461
  "f(x:=y)"                       == "CONST update(f,x,y)"
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   462
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   463
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   464
lemma update_apply [simp]: "f(x:=y) ` z = (if z=x then y else f`z)"
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   465
apply (simp add: update_def)
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   466
apply (case_tac "z \<in> domain(f)")
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   467
apply (simp_all add: apply_0)
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   468
done
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   469
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   470
lemma update_idem: "\<lbrakk>f`x = y;  f \<in> Pi(A,B);  x \<in> A\<rbrakk> \<Longrightarrow> f(x:=y) = f"
76216
9fc34f76b4e8 getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents: 76215
diff changeset
   471
  unfolding update_def
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   472
apply (simp add: domain_of_fun cons_absorb)
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   473
apply (rule fun_extension)
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   474
apply (best intro: apply_type if_type lam_type, assumption, simp)
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   475
done
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   476
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   477
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   478
(* \<lbrakk>f \<in> Pi(A, B); x \<in> A\<rbrakk> \<Longrightarrow> f(x := f`x) = f *)
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   479
declare refl [THEN update_idem, simp]
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   480
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   481
lemma domain_update [simp]: "domain(f(x:=y)) = cons(x, domain(f))"
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   482
by (unfold update_def, simp)
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   483
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   484
lemma update_type: "\<lbrakk>f \<in> Pi(A,B);  x \<in> A;  y \<in> B(x)\<rbrakk> \<Longrightarrow> f(x:=y) \<in> Pi(A, B)"
76216
9fc34f76b4e8 getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents: 76215
diff changeset
   485
  unfolding update_def
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   486
apply (simp add: domain_of_fun cons_absorb apply_funtype lam_type)
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   487
done
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   488
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   489
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   490
subsection\<open>Monotonicity Theorems\<close>
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   491
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   492
subsubsection\<open>Replacement in its Various Forms\<close>
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   493
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   494
(*Not easy to express monotonicity in P, since any "bigger" predicate
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   495
  would have to be single-valued*)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   496
lemma Replace_mono: "A<=B \<Longrightarrow> Replace(A,P) \<subseteq> Replace(B,P)"
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   497
by (blast elim!: ReplaceE)
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   498
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   499
lemma RepFun_mono: "A<=B \<Longrightarrow> {f(x). x \<in> A} \<subseteq> {f(x). x \<in> B}"
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   500
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   501
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   502
lemma Pow_mono: "A<=B \<Longrightarrow> Pow(A) \<subseteq> Pow(B)"
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   503
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   504
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   505
lemma Union_mono: "A<=B \<Longrightarrow> \<Union>(A) \<subseteq> \<Union>(B)"
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   506
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   507
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   508
lemma UN_mono:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   509
    "\<lbrakk>A<=C;  \<And>x. x \<in> A \<Longrightarrow> B(x)<=D(x)\<rbrakk> \<Longrightarrow> (\<Union>x\<in>A. B(x)) \<subseteq> (\<Union>x\<in>C. D(x))"
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   510
by blast
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   511
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   512
(*Intersection is ANTI-monotonic.  There are TWO premises! *)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   513
lemma Inter_anti_mono: "\<lbrakk>A<=B;  A\<noteq>0\<rbrakk> \<Longrightarrow> \<Inter>(B) \<subseteq> \<Inter>(A)"
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   514
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   515
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   516
lemma cons_mono: "C<=D \<Longrightarrow> cons(a,C) \<subseteq> cons(a,D)"
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   517
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   518
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   519
lemma Un_mono: "\<lbrakk>A<=C;  B<=D\<rbrakk> \<Longrightarrow> A \<union> B \<subseteq> C \<union> D"
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   520
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   521
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   522
lemma Int_mono: "\<lbrakk>A<=C;  B<=D\<rbrakk> \<Longrightarrow> A \<inter> B \<subseteq> C \<inter> D"
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   523
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   524
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   525
lemma Diff_mono: "\<lbrakk>A<=C;  D<=B\<rbrakk> \<Longrightarrow> A-B \<subseteq> C-D"
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   526
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   527
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   528
subsubsection\<open>Standard Products, Sums and Function Spaces\<close>
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   529
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   530
lemma Sigma_mono [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   531
     "\<lbrakk>A<=C;  \<And>x. x \<in> A \<longrightarrow> B(x) \<subseteq> D(x)\<rbrakk> \<Longrightarrow> Sigma(A,B) \<subseteq> Sigma(C,D)"
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   532
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   533
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   534
lemma sum_mono: "\<lbrakk>A<=C;  B<=D\<rbrakk> \<Longrightarrow> A+B \<subseteq> C+D"
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   535
by (unfold sum_def, blast)
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   536
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   537
(*Note that B->A and C->A are typically disjoint!*)
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   538
lemma Pi_mono: "B<=C \<Longrightarrow> A->B \<subseteq> A->C"
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   539
by (blast intro: lam_type elim: Pi_lamE)
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   540
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   541
lemma lam_mono: "A<=B \<Longrightarrow> Lambda(A,c) \<subseteq> Lambda(B,c)"
76216
9fc34f76b4e8 getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents: 76215
diff changeset
   542
  unfolding lam_def
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   543
apply (erule RepFun_mono)
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   544
done
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   545
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   546
subsubsection\<open>Converse, Domain, Range, Field\<close>
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   547
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   548
lemma converse_mono: "r<=s \<Longrightarrow> converse(r) \<subseteq> converse(s)"
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   549
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   550
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   551
lemma domain_mono: "r<=s \<Longrightarrow> domain(r)<=domain(s)"
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   552
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   553
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   554
lemmas domain_rel_subset = subset_trans [OF domain_mono domain_subset]
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   555
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   556
lemma range_mono: "r<=s \<Longrightarrow> range(r)<=range(s)"
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   557
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   558
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   559
lemmas range_rel_subset = subset_trans [OF range_mono range_subset]
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   560
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   561
lemma field_mono: "r<=s \<Longrightarrow> field(r)<=field(s)"
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   562
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   563
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   564
lemma field_rel_subset: "r \<subseteq> A*A \<Longrightarrow> field(r) \<subseteq> A"
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   565
by (erule field_mono [THEN subset_trans], blast)
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   566
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   567
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   568
subsubsection\<open>Images\<close>
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   569
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   570
lemma image_pair_mono:
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   571
    "\<lbrakk>\<And>x y. \<langle>x,y\<rangle>:r \<Longrightarrow> \<langle>x,y\<rangle>:s;  A<=B\<rbrakk> \<Longrightarrow> r``A \<subseteq> s``B"
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   572
by blast
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   573
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   574
lemma vimage_pair_mono:
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   575
    "\<lbrakk>\<And>x y. \<langle>x,y\<rangle>:r \<Longrightarrow> \<langle>x,y\<rangle>:s;  A<=B\<rbrakk> \<Longrightarrow> r-``A \<subseteq> s-``B"
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   576
by blast
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   577
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   578
lemma image_mono: "\<lbrakk>r<=s;  A<=B\<rbrakk> \<Longrightarrow> r``A \<subseteq> s``B"
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   579
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   580
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   581
lemma vimage_mono: "\<lbrakk>r<=s;  A<=B\<rbrakk> \<Longrightarrow> r-``A \<subseteq> s-``B"
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   582
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   583
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   584
lemma Collect_mono:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   585
    "\<lbrakk>A<=B;  \<And>x. x \<in> A \<Longrightarrow> P(x) \<longrightarrow> Q(x)\<rbrakk> \<Longrightarrow> Collect(A,P) \<subseteq> Collect(B,Q)"
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   586
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   587
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   588
(*Used in intr_elim.ML and in individual datatype definitions*)
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   589
lemmas basic_monos = subset_refl imp_refl disj_mono conj_mono ex_mono
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   590
                     Collect_mono Part_mono in_mono
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   591
27702
80608e96e760 Lemmas added
ballarin
parents: 27150
diff changeset
   592
(* Useful with simp; contributed by Clemens Ballarin. *)
80608e96e760 Lemmas added
ballarin
parents: 27150
diff changeset
   593
80608e96e760 Lemmas added
ballarin
parents: 27150
diff changeset
   594
lemma bex_image_simp:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   595
  "\<lbrakk>f \<in> Pi(X, Y); A \<subseteq> X\<rbrakk>  \<Longrightarrow> (\<exists>x\<in>f``A. P(x)) \<longleftrightarrow> (\<exists>x\<in>A. P(f`x))"
27702
80608e96e760 Lemmas added
ballarin
parents: 27150
diff changeset
   596
  apply safe
80608e96e760 Lemmas added
ballarin
parents: 27150
diff changeset
   597
   apply rule
80608e96e760 Lemmas added
ballarin
parents: 27150
diff changeset
   598
    prefer 2 apply assumption
80608e96e760 Lemmas added
ballarin
parents: 27150
diff changeset
   599
   apply (simp add: apply_equality)
80608e96e760 Lemmas added
ballarin
parents: 27150
diff changeset
   600
  apply (blast intro: apply_Pair)
80608e96e760 Lemmas added
ballarin
parents: 27150
diff changeset
   601
  done
80608e96e760 Lemmas added
ballarin
parents: 27150
diff changeset
   602
80608e96e760 Lemmas added
ballarin
parents: 27150
diff changeset
   603
lemma ball_image_simp:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 71085
diff changeset
   604
  "\<lbrakk>f \<in> Pi(X, Y); A \<subseteq> X\<rbrakk>  \<Longrightarrow> (\<forall>x\<in>f``A. P(x)) \<longleftrightarrow> (\<forall>x\<in>A. P(f`x))"
27702
80608e96e760 Lemmas added
ballarin
parents: 27150
diff changeset
   605
  apply safe
80608e96e760 Lemmas added
ballarin
parents: 27150
diff changeset
   606
   apply (blast intro: apply_Pair)
80608e96e760 Lemmas added
ballarin
parents: 27150
diff changeset
   607
  apply (drule bspec) apply assumption
80608e96e760 Lemmas added
ballarin
parents: 27150
diff changeset
   608
  apply (simp add: apply_equality)
80608e96e760 Lemmas added
ballarin
parents: 27150
diff changeset
   609
  done
80608e96e760 Lemmas added
ballarin
parents: 27150
diff changeset
   610
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   611
end