src/ZF/func.thy
author wenzelm
Fri, 17 Nov 2006 02:20:03 +0100
changeset 21404 eb85850d3eb7
parent 16417 9bc16273c2d4
child 24893 b8ef7afe3a6b
permissions -rw-r--r--
more robust syntax for definition/abbreviation/notation;
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
    ID:         $Id$
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
     4
    Copyright   1991  University of Cambridge
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
     5
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
     6
*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
     7
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
     8
header{*Functions, Function Spaces, Lambda-Abstraction*}
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
     9
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14883
diff changeset
    10
theory func imports equalities Sum begin
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    11
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
    12
subsection{*The Pi Operator: Dependent Function Space*}
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
    13
13248
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
    14
lemma subset_Sigma_imp_relation: "r <= Sigma(A,B) ==> relation(r)"
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
    15
by (simp add: relation_def, blast)
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
    16
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 13219
diff changeset
    17
lemma relation_converse_converse [simp]:
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 13219
diff changeset
    18
     "relation(r) ==> converse(converse(r)) = r"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 13219
diff changeset
    19
by (simp add: relation_def, blast) 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 13219
diff changeset
    20
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 13219
diff changeset
    21
lemma relation_restrict [simp]:  "relation(restrict(r,A))"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 13219
diff changeset
    22
by (simp add: restrict_def relation_def, blast) 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 13219
diff changeset
    23
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    24
lemma Pi_iff:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    25
    "f: Pi(A,B) <-> function(f) & f<=Sigma(A,B) & A<=domain(f)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    26
by (unfold Pi_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    27
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    28
(*For upward compatibility with the former definition*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    29
lemma Pi_iff_old:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    30
    "f: Pi(A,B) <-> f<=Sigma(A,B) & (ALL x:A. EX! y. <x,y>: f)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    31
by (unfold Pi_def function_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    32
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    33
lemma fun_is_function: "f: Pi(A,B) ==> function(f)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    34
by (simp only: Pi_iff)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    35
13219
7e44aa8a276e new lemma
paulson
parents: 13179
diff changeset
    36
lemma function_imp_Pi:
7e44aa8a276e new lemma
paulson
parents: 13179
diff changeset
    37
     "[|function(f); relation(f)|] ==> f \<in> domain(f) -> range(f)"
7e44aa8a276e new lemma
paulson
parents: 13179
diff changeset
    38
by (simp add: Pi_iff relation_def, blast) 
7e44aa8a276e new lemma
paulson
parents: 13179
diff changeset
    39
13172
03a5afa7b888 tidying up
paulson
parents: 13168
diff changeset
    40
lemma functionI: 
03a5afa7b888 tidying up
paulson
parents: 13168
diff changeset
    41
     "[| !!x y y'. [| <x,y>:r; <x,y'>:r |] ==> y=y' |] ==> function(r)"
03a5afa7b888 tidying up
paulson
parents: 13168
diff changeset
    42
by (simp add: function_def, blast) 
03a5afa7b888 tidying up
paulson
parents: 13168
diff changeset
    43
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    44
(*Functions are relations*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    45
lemma fun_is_rel: "f: Pi(A,B) ==> f <= Sigma(A,B)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    46
by (unfold Pi_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    47
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    48
lemma Pi_cong:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    49
    "[| A=A';  !!x. x:A' ==> B(x)=B'(x) |] ==> Pi(A,B) = Pi(A',B')"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    50
by (simp add: Pi_def cong add: Sigma_cong)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    51
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    52
(*Sigma_cong, Pi_cong NOT given to Addcongs: they cause
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    53
  flex-flex pairs and the "Check your prover" error.  Most
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    54
  Sigmas and Pis are abbreviated as * or -> *)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    55
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    56
(*Weakening one function type to another; see also Pi_type*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    57
lemma fun_weaken_type: "[| f: A->B;  B<=D |] ==> f: A->D"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    58
by (unfold Pi_def, best)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    59
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
    60
subsection{*Function Application*}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    61
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    62
lemma apply_equality2: "[| <a,b>: f;  <a,c>: f;  f: Pi(A,B) |] ==> b=c"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    63
by (unfold Pi_def function_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    64
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    65
lemma function_apply_equality: "[| <a,b>: f;  function(f) |] ==> f`a = b"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    66
by (unfold apply_def function_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    67
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    68
lemma apply_equality: "[| <a,b>: f;  f: Pi(A,B) |] ==> f`a = b"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    69
apply (unfold Pi_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    70
apply (blast intro: function_apply_equality)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    71
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    72
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    73
(*Applying a function outside its domain yields 0*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    74
lemma apply_0: "a ~: domain(f) ==> f`a = 0"
13176
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
    75
by (unfold apply_def, blast)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    76
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    77
lemma Pi_memberD: "[| f: Pi(A,B);  c: f |] ==> EX x:A.  c = <x,f`x>"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    78
apply (frule fun_is_rel)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    79
apply (blast dest: apply_equality)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    80
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    81
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    82
lemma function_apply_Pair: "[| function(f);  a : domain(f)|] ==> <a,f`a>: f"
13176
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
    83
apply (simp add: function_def, clarify) 
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
    84
apply (subgoal_tac "f`a = y", blast) 
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
    85
apply (simp add: apply_def, blast) 
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    86
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    87
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    88
lemma apply_Pair: "[| f: Pi(A,B);  a:A |] ==> <a,f`a>: f"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    89
apply (simp add: Pi_iff)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    90
apply (blast intro: function_apply_Pair)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    91
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    92
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    93
(*Conclusion is flexible -- use res_inst_tac or else apply_funtype below!*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    94
lemma apply_type [TC]: "[| f: Pi(A,B);  a:A |] ==> f`a : B(a)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    95
by (blast intro: apply_Pair dest: fun_is_rel)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    96
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    97
(*This version is acceptable to the simplifier*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    98
lemma apply_funtype: "[| f: A->B;  a:A |] ==> f`a : B"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
    99
by (blast dest: apply_type)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   100
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   101
lemma apply_iff: "f: Pi(A,B) ==> <a,b>: f <-> a:A & f`a = b"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   102
apply (frule fun_is_rel)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   103
apply (blast intro!: apply_Pair apply_equality)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   104
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   105
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   106
(*Refining one Pi type to another*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   107
lemma Pi_type: "[| f: Pi(A,C);  !!x. x:A ==> f`x : B(x) |] ==> f : Pi(A,B)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   108
apply (simp only: Pi_iff)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   109
apply (blast dest: function_apply_equality)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   110
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   111
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   112
(*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
   113
lemma Pi_Collect_iff:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   114
     "(f : Pi(A, %x. {y:B(x). P(x,y)}))
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   115
      <->  f : Pi(A,B) & (ALL x: A. P(x, f`x))"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   116
by (blast intro: Pi_type dest: apply_type)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   117
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   118
lemma Pi_weaken_type:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   119
        "[| f : Pi(A,B);  !!x. x:A ==> B(x)<=C(x) |] ==> f : Pi(A,C)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   120
by (blast intro: Pi_type dest: apply_type)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   121
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   122
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   123
(** Elimination of membership in a function **)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   124
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   125
lemma domain_type: "[| <a,b> : f;  f: Pi(A,B) |] ==> a : A"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   126
by (blast dest: fun_is_rel)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   127
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   128
lemma range_type: "[| <a,b> : f;  f: Pi(A,B) |] ==> b : B(a)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   129
by (blast dest: fun_is_rel)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   130
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   131
lemma Pair_mem_PiD: "[| <a,b>: f;  f: Pi(A,B) |] ==> a:A & b:B(a) & f`a = b"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   132
by (blast intro: domain_type range_type apply_equality)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   133
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   134
subsection{*Lambda Abstraction*}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   135
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   136
lemma lamI: "a:A ==> <a,b(a)> : (lam x:A. b(x))"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   137
apply (unfold lam_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   138
apply (erule RepFunI)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   139
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   140
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   141
lemma lamE:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   142
    "[| p: (lam x:A. b(x));  !!x.[| x:A; p=<x,b(x)> |] ==> P
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   143
     |] ==>  P"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   144
by (simp add: lam_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   145
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   146
lemma lamD: "[| <a,c>: (lam x:A. b(x)) |] ==> c = b(a)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   147
by (simp add: lam_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   148
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   149
lemma lam_type [TC]:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   150
    "[| !!x. x:A ==> b(x): B(x) |] ==> (lam x:A. b(x)) : Pi(A,B)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   151
by (simp add: lam_def Pi_def function_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   152
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   153
lemma lam_funtype: "(lam x:A. b(x)) : A -> {b(x). x:A}"
13176
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
   154
by (blast intro: lam_type)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   155
13172
03a5afa7b888 tidying up
paulson
parents: 13168
diff changeset
   156
lemma function_lam: "function (lam x:A. b(x))"
03a5afa7b888 tidying up
paulson
parents: 13168
diff changeset
   157
by (simp add: function_def lam_def) 
03a5afa7b888 tidying up
paulson
parents: 13168
diff changeset
   158
03a5afa7b888 tidying up
paulson
parents: 13168
diff changeset
   159
lemma relation_lam: "relation (lam x:A. b(x))"  
03a5afa7b888 tidying up
paulson
parents: 13168
diff changeset
   160
by (simp add: relation_def lam_def) 
03a5afa7b888 tidying up
paulson
parents: 13168
diff changeset
   161
13175
81082cfa5618 new definition of "apply" and new simprule "beta_if"
paulson
parents: 13174
diff changeset
   162
lemma beta_if [simp]: "(lam x:A. b(x)) ` a = (if a : A then b(a) else 0)"
13176
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
   163
by (simp add: apply_def lam_def, blast)
13175
81082cfa5618 new definition of "apply" and new simprule "beta_if"
paulson
parents: 13174
diff changeset
   164
81082cfa5618 new definition of "apply" and new simprule "beta_if"
paulson
parents: 13174
diff changeset
   165
lemma beta: "a : A ==> (lam x:A. b(x)) ` a = b(a)"
13176
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
   166
by (simp add: apply_def lam_def, blast)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   167
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   168
lemma lam_empty [simp]: "(lam x:0. b(x)) = 0"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   169
by (simp add: lam_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   170
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   171
lemma domain_lam [simp]: "domain(Lambda(A,b)) = A"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   172
by (simp add: lam_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   173
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   174
(*congruence rule for lambda abstraction*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   175
lemma lam_cong [cong]:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   176
    "[| A=A';  !!x. x:A' ==> b(x)=b'(x) |] ==> Lambda(A,b) = Lambda(A',b')"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   177
by (simp only: lam_def cong add: RepFun_cong)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   178
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   179
lemma lam_theI:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   180
    "(!!x. x:A ==> EX! y. Q(x,y)) ==> EX f. ALL x:A. Q(x, f`x)"
13175
81082cfa5618 new definition of "apply" and new simprule "beta_if"
paulson
parents: 13174
diff changeset
   181
apply (rule_tac x = "lam x: A. THE y. Q (x,y)" in exI)
13176
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
   182
apply simp 
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   183
apply (blast intro: theI)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   184
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   185
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   186
lemma lam_eqE: "[| (lam x:A. f(x)) = (lam x:A. g(x));  a:A |] ==> f(a)=g(a)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   187
by (fast intro!: lamI elim: equalityE lamE)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   188
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   189
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   190
(*Empty function spaces*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   191
lemma Pi_empty1 [simp]: "Pi(0,A) = {0}"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   192
by (unfold Pi_def function_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   193
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   194
(*The singleton function*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   195
lemma singleton_fun [simp]: "{<a,b>} : {a} -> {b}"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   196
by (unfold Pi_def function_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   197
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   198
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
   199
by (unfold Pi_def function_def, force)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   200
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   201
lemma  fun_space_empty_iff [iff]: "(A->X)=0 \<longleftrightarrow> X=0 & (A \<noteq> 0)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   202
apply auto
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   203
apply (fast intro!: equals0I intro: lam_type)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   204
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   205
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   206
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   207
subsection{*Extensionality*}
13163
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
(*Semi-extensionality!*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   210
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   211
lemma fun_subset:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   212
    "[| f : Pi(A,B);  g: Pi(C,D);  A<=C;
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   213
        !!x. x:A ==> f`x = g`x       |] ==> f<=g"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   214
by (force dest: Pi_memberD intro: apply_Pair)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   215
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   216
lemma fun_extension:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   217
    "[| f : Pi(A,B);  g: Pi(A,D);
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   218
        !!x. x:A ==> f`x = g`x       |] ==> f=g"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   219
by (blast del: subsetI intro: subset_refl sym fun_subset)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   220
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   221
lemma eta [simp]: "f : Pi(A,B) ==> (lam x:A. f`x) = f"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   222
apply (rule fun_extension)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   223
apply (auto simp add: lam_type apply_type beta)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   224
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   225
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   226
lemma fun_extension_iff:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   227
     "[| f:Pi(A,B); g:Pi(A,C) |] ==> (ALL a:A. f`a = g`a) <-> f=g"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   228
by (blast intro: fun_extension)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   229
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   230
(*thm by Mark Staples, proof by lcp*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   231
lemma fun_subset_eq: "[| f:Pi(A,B); g:Pi(A,C) |] ==> f <= g <-> (f = g)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   232
by (blast dest: apply_Pair
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   233
	  intro: fun_extension apply_equality [symmetric])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   234
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   235
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   236
(*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
   237
lemma Pi_lamE:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   238
  assumes major: "f: Pi(A,B)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   239
      and minor: "!!b. [| ALL x:A. b(x):B(x);  f = (lam x:A. b(x)) |] ==> P"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   240
  shows "P"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   241
apply (rule minor)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   242
apply (rule_tac [2] eta [symmetric])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   243
apply (blast intro: major apply_type)+
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   244
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   245
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   246
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   247
subsection{*Images of Functions*}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   248
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   249
lemma image_lam: "C <= A ==> (lam x:A. b(x)) `` C = {b(x). x:C}"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   250
by (unfold lam_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   251
13179
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   252
lemma Repfun_function_if:
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   253
     "function(f) 
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   254
      ==> {f`x. x:C} = (if C <= domain(f) then f``C else cons(0,f``C))";
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   255
apply simp
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   256
apply (intro conjI impI)  
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   257
 apply (blast dest: function_apply_equality intro: function_apply_Pair) 
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   258
apply (rule equalityI)
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   259
 apply (blast intro!: function_apply_Pair apply_0) 
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   260
apply (blast dest: function_apply_equality intro: apply_0 [symmetric]) 
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   261
done
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   262
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   263
(*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
   264
  be written \<Union>x\<in>C. {f`x} *)
13174
85d3c0981a16 more tidying
paulson
parents: 13172
diff changeset
   265
lemma image_function:
85d3c0981a16 more tidying
paulson
parents: 13172
diff changeset
   266
     "[| function(f);  C <= domain(f) |] ==> f``C = {f`x. x:C}";
13179
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   267
by (simp add: Repfun_function_if) 
13174
85d3c0981a16 more tidying
paulson
parents: 13172
diff changeset
   268
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   269
lemma image_fun: "[| f : Pi(A,B);  C <= A |] ==> f``C = {f`x. x:C}"
13174
85d3c0981a16 more tidying
paulson
parents: 13172
diff changeset
   270
apply (simp add: Pi_iff) 
85d3c0981a16 more tidying
paulson
parents: 13172
diff changeset
   271
apply (blast intro: image_function) 
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   272
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   273
14883
ca000a495448 Groups, Rings and supporting lemmas
paulson
parents: 14153
diff changeset
   274
lemma image_eq_UN: 
ca000a495448 Groups, Rings and supporting lemmas
paulson
parents: 14153
diff changeset
   275
  assumes f: "f \<in> Pi(A,B)" "C \<subseteq> A" shows "f``C = (\<Union>x\<in>C. {f ` x})"
ca000a495448 Groups, Rings and supporting lemmas
paulson
parents: 14153
diff changeset
   276
by (auto simp add: image_fun [OF f]) 
ca000a495448 Groups, Rings and supporting lemmas
paulson
parents: 14153
diff changeset
   277
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   278
lemma Pi_image_cons:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   279
     "[| f: Pi(A,B);  x: A |] ==> f `` cons(x,y) = cons(f`x, f``y)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   280
by (blast dest: apply_equality apply_Pair)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   281
124
858ab9a9b047 made pseudo theories for all ML files;
clasohm
parents:
diff changeset
   282
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   283
subsection{*Properties of @{term "restrict(f,A)"}*}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   284
13179
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   285
lemma restrict_subset: "restrict(f,A) <= f"
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   286
by (unfold restrict_def, blast)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   287
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   288
lemma function_restrictI:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   289
    "function(f) ==> function(restrict(f,A))"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   290
by (unfold restrict_def function_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   291
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   292
lemma restrict_type2: "[| f: Pi(C,B);  A<=C |] ==> restrict(f,A) : Pi(A,B)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   293
by (simp add: Pi_iff function_def restrict_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   294
13179
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   295
lemma restrict: "restrict(f,A) ` a = (if a : A then f`a else 0)"
13176
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
   296
by (simp add: apply_def restrict_def, blast)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   297
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   298
lemma restrict_empty [simp]: "restrict(f,0) = 0"
13179
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   299
by (unfold restrict_def, simp)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   300
13172
03a5afa7b888 tidying up
paulson
parents: 13168
diff changeset
   301
lemma restrict_iff: "z \<in> restrict(r,A) \<longleftrightarrow> z \<in> r & (\<exists>x\<in>A. \<exists>y. z = \<langle>x, y\<rangle>)"
03a5afa7b888 tidying up
paulson
parents: 13168
diff changeset
   302
by (simp add: restrict_def) 
03a5afa7b888 tidying up
paulson
parents: 13168
diff changeset
   303
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   304
lemma restrict_restrict [simp]:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   305
     "restrict(restrict(r,A),B) = restrict(r, A Int B)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   306
by (unfold restrict_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   307
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   308
lemma domain_restrict [simp]: "domain(restrict(f,C)) = domain(f) Int C"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   309
apply (unfold restrict_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   310
apply (auto simp add: domain_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   311
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   312
13248
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   313
lemma restrict_idem: "f <= Sigma(A,B) ==> restrict(f,A) = f"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   314
by (simp add: restrict_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   315
13248
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   316
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   317
(*converse probably holds too*)
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   318
lemma domain_restrict_idem:
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   319
     "[| domain(r) <= A; relation(r) |] ==> restrict(r,A) = r"
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   320
by (simp add: restrict_def relation_def, blast)
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   321
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   322
lemma domain_restrict_lam [simp]: "domain(restrict(Lambda(A,f),C)) = A Int C"
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   323
apply (unfold restrict_def lam_def)
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   324
apply (rule equalityI)
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   325
apply (auto simp add: domain_iff)
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   326
done
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   327
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   328
lemma restrict_if [simp]: "restrict(f,A) ` a = (if a : A then f`a else 0)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   329
by (simp add: restrict apply_0)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   330
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   331
lemma restrict_lam_eq:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   332
    "A<=C ==> restrict(lam x:C. b(x), A) = (lam x:A. b(x))"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   333
by (unfold restrict_def lam_def, auto)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   334
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   335
lemma fun_cons_restrict_eq:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   336
     "f : cons(a, b) -> B ==> f = cons(<a, f ` a>, restrict(f, b))"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   337
apply (rule equalityI)
13248
ae66c22ed52e new theorems
paulson
parents: 13221
diff changeset
   338
 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
   339
apply (auto dest!: Pi_memberD simp add: restrict_def lam_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   340
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   341
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   342
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   343
subsection{*Unions of Functions*}
13163
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
(** The Union of a set of COMPATIBLE functions is a function **)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   346
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   347
lemma function_Union:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   348
    "[| ALL x:S. function(x);
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   349
        ALL x:S. ALL y:S. x<=y | y<=x  |]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   350
     ==> function(Union(S))"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   351
by (unfold function_def, blast) 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   352
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   353
lemma fun_Union:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   354
    "[| ALL f:S. EX C D. f:C->D;
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   355
             ALL f:S. ALL y:S. f<=y | y<=f  |] ==>
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   356
          Union(S) : domain(Union(S)) -> range(Union(S))"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   357
apply (unfold Pi_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   358
apply (blast intro!: rel_Union function_Union)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   359
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   360
13174
85d3c0981a16 more tidying
paulson
parents: 13172
diff changeset
   361
lemma gen_relation_Union [rule_format]:
85d3c0981a16 more tidying
paulson
parents: 13172
diff changeset
   362
     "\<forall>f\<in>F. relation(f) \<Longrightarrow> relation(Union(F))"
85d3c0981a16 more tidying
paulson
parents: 13172
diff changeset
   363
by (simp add: relation_def) 
85d3c0981a16 more tidying
paulson
parents: 13172
diff changeset
   364
13163
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
(** The Union of 2 disjoint functions is a function **)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   367
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   368
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
   369
                subset_trans [OF _ Un_upper1]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   370
                subset_trans [OF _ Un_upper2]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   371
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   372
lemma fun_disjoint_Un:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   373
     "[| f: A->B;  g: C->D;  A Int C = 0  |]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   374
      ==> (f Un g) : (A Un C) -> (B Un D)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   375
(*Prove the product and domain subgoals using distributive laws*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   376
apply (simp add: Pi_iff extension Un_rls)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   377
apply (unfold function_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   378
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   379
13179
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   380
lemma fun_disjoint_apply1: "a \<notin> domain(g) ==> (f Un g)`a = f`a"
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   381
by (simp add: apply_def, blast) 
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   382
13179
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   383
lemma fun_disjoint_apply2: "c \<notin> domain(f) ==> (f Un g)`c = g`c"
3f6f00c6c56f strong lemmas about functions
paulson
parents: 13176
diff changeset
   384
by (simp add: apply_def, blast) 
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   385
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   386
subsection{*Domain and Range of a Function or Relation*}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   387
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   388
lemma domain_of_fun: "f : Pi(A,B) ==> domain(f)=A"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   389
by (unfold Pi_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   390
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   391
lemma apply_rangeI: "[| f : Pi(A,B);  a: A |] ==> f`a : range(f)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   392
by (erule apply_Pair [THEN rangeI], assumption)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   393
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   394
lemma range_of_fun: "f : Pi(A,B) ==> f : A->range(f)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   395
by (blast intro: Pi_type apply_rangeI)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   396
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   397
subsection{*Extensions of Functions*}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   398
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   399
lemma fun_extend:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   400
     "[| f: A->B;  c~:A |] ==> cons(<c,b>,f) : cons(c,A) -> cons(b,B)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   401
apply (frule singleton_fun [THEN fun_disjoint_Un], blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   402
apply (simp add: cons_eq) 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   403
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   404
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   405
lemma fun_extend3:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   406
     "[| f: A->B;  c~:A;  b: B |] ==> cons(<c,b>,f) : cons(c,A) -> B"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   407
by (blast intro: fun_extend [THEN fun_weaken_type])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   408
13176
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
   409
lemma extend_apply:
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
   410
     "c ~: domain(f) ==> cons(<c,b>,f)`a = (if a=c then b else f`a)"
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
   411
by (auto simp add: apply_def) 
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   412
13176
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
   413
lemma fun_extend_apply [simp]:
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
   414
     "[| f: A->B;  c~:A |] ==> cons(<c,b>,f)`a = (if a=c then b else f`a)" 
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
   415
apply (rule extend_apply) 
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
   416
apply (simp add: Pi_def, blast) 
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   417
done
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
lemmas singleton_apply = apply_equality [OF singletonI singleton_fun, simp]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   420
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   421
(*For Finite.ML.  Inclusion of right into left is easy*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   422
lemma cons_fun_eq:
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13248
diff changeset
   423
     "c ~: A ==> cons(c,A) -> B = (\<Union>f \<in> A->B. \<Union>b\<in>B. {cons(<c,b>, f)})"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   424
apply (rule equalityI)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   425
apply (safe elim!: fun_extend3)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   426
(*Inclusion of left into right*)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   427
apply (subgoal_tac "restrict (x, A) : A -> B")
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   428
 prefer 2 apply (blast intro: restrict_type2)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   429
apply (rule UN_I, assumption)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   430
apply (rule apply_funtype [THEN UN_I]) 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   431
  apply assumption
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   432
 apply (rule consI1) 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   433
apply (simp (no_asm))
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   434
apply (rule fun_extension) 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   435
  apply assumption
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   436
 apply (blast intro: fun_extend) 
13176
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
   437
apply (erule consE, simp_all)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   438
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   439
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13248
diff changeset
   440
lemma succ_fun_eq: "succ(n) -> B = (\<Union>f \<in> n->B. \<Union>b\<in>B. {cons(<n,b>, f)})"
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13248
diff changeset
   441
by (simp add: succ_def mem_not_refl cons_fun_eq)
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13248
diff changeset
   442
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   443
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   444
subsection{*Function Updates*}
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   445
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   446
constdefs
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   447
  update  :: "[i,i,i] => i"
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   448
   "update(f,a,b) == lam x: cons(a, domain(f)). if(x=a, b, f`x)"
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   449
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   450
nonterminals
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   451
  updbinds  updbind
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   452
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   453
syntax
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   454
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   455
  (* Let expressions *)
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   456
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   457
  "_updbind"    :: "[i, i] => updbind"               ("(2_ :=/ _)")
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   458
  ""            :: "updbind => updbinds"             ("_")
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   459
  "_updbinds"   :: "[updbind, updbinds] => updbinds" ("_,/ _")
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   460
  "_Update"     :: "[i, updbinds] => i"              ("_/'((_)')" [900,0] 900)
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   461
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   462
translations
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   463
  "_Update (f, _updbinds(b,bs))"  == "_Update (_Update(f,b), bs)"
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   464
  "f(x:=y)"                       == "update(f,x,y)"
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   465
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   466
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   467
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
   468
apply (simp add: update_def)
14153
76a6ba67bd15 new case_tac
paulson
parents: 14095
diff changeset
   469
apply (case_tac "z \<in> domain(f)")   
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   470
apply (simp_all add: apply_0)
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   471
done
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   472
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   473
lemma update_idem: "[| f`x = y;  f: Pi(A,B);  x: A |] ==> f(x:=y) = f"
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   474
apply (unfold update_def)
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   475
apply (simp add: domain_of_fun cons_absorb)
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   476
apply (rule fun_extension)
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   477
apply (best intro: apply_type if_type lam_type, assumption, simp)
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   478
done
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   479
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   480
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   481
(* [| f: Pi(A, B); x:A |] ==> f(x := f`x) = f *)
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   482
declare refl [THEN update_idem, simp]
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   483
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   484
lemma domain_update [simp]: "domain(f(x:=y)) = cons(x, domain(f))"
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   485
by (unfold update_def, simp)
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   486
14046
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 13615
diff changeset
   487
lemma update_type: "[| f:Pi(A,B);  x : A;  y: B(x) |] ==> f(x:=y) : Pi(A, B)"
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   488
apply (unfold update_def)
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   489
apply (simp add: domain_of_fun cons_absorb apply_funtype lam_type)
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   490
done
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   491
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   492
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   493
subsection{*Monotonicity Theorems*}
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   494
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   495
subsubsection{*Replacement in its Various Forms*}
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   496
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   497
(*Not easy to express monotonicity in P, since any "bigger" predicate
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   498
  would have to be single-valued*)
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   499
lemma Replace_mono: "A<=B ==> Replace(A,P) <= Replace(B,P)"
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   500
by (blast elim!: ReplaceE)
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   501
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   502
lemma RepFun_mono: "A<=B ==> {f(x). x:A} <= {f(x). x:B}"
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   503
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   504
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   505
lemma Pow_mono: "A<=B ==> Pow(A) <= Pow(B)"
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 Union_mono: "A<=B ==> Union(A) <= Union(B)"
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   509
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   510
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   511
lemma UN_mono:
13615
449a70d88b38 Numerous cosmetic changes, prompted by the new simplifier
paulson
parents: 13357
diff changeset
   512
    "[| A<=C;  !!x. x:A ==> B(x)<=D(x) |] ==> (\<Union>x\<in>A. B(x)) <= (\<Union>x\<in>C. D(x))"
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   513
by blast  
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   514
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   515
(*Intersection is ANTI-monotonic.  There are TWO premises! *)
14095
a1ba833d6b61 Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
paulson
parents: 14046
diff changeset
   516
lemma Inter_anti_mono: "[| A<=B;  A\<noteq>0 |] ==> Inter(B) <= Inter(A)"
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
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   519
lemma cons_mono: "C<=D ==> cons(a,C) <= cons(a,D)"
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   520
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   521
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   522
lemma Un_mono: "[| A<=C;  B<=D |] ==> A Un B <= C Un D"
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   523
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   524
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   525
lemma Int_mono: "[| A<=C;  B<=D |] ==> A Int B <= C Int D"
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   526
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   527
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   528
lemma Diff_mono: "[| A<=C;  D<=B |] ==> A-B <= C-D"
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   529
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   530
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   531
subsubsection{*Standard Products, Sums and Function Spaces *}
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   532
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   533
lemma Sigma_mono [rule_format]:
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   534
     "[| A<=C;  !!x. x:A --> B(x) <= D(x) |] ==> Sigma(A,B) <= Sigma(C,D)"
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   535
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   536
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   537
lemma sum_mono: "[| A<=C;  B<=D |] ==> A+B <= C+D"
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   538
by (unfold sum_def, blast)
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   539
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   540
(*Note that B->A and C->A are typically disjoint!*)
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   541
lemma Pi_mono: "B<=C ==> A->B <= A->C"
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   542
by (blast intro: lam_type elim: Pi_lamE)
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   543
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   544
lemma lam_mono: "A<=B ==> Lambda(A,c) <= Lambda(B,c)"
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   545
apply (unfold lam_def)
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   546
apply (erule RepFun_mono)
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   547
done
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   548
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   549
subsubsection{*Converse, Domain, Range, Field*}
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   550
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   551
lemma converse_mono: "r<=s ==> converse(r) <= converse(s)"
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
lemma domain_mono: "r<=s ==> domain(r)<=domain(s)"
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   555
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   556
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   557
lemmas domain_rel_subset = subset_trans [OF domain_mono domain_subset]
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   558
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   559
lemma range_mono: "r<=s ==> range(r)<=range(s)"
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   560
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   561
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   562
lemmas range_rel_subset = subset_trans [OF range_mono range_subset]
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   563
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   564
lemma field_mono: "r<=s ==> field(r)<=field(s)"
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   565
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   566
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   567
lemma field_rel_subset: "r <= A*A ==> field(r) <= A"
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   568
by (erule field_mono [THEN subset_trans], blast)
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   569
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   570
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   571
subsubsection{*Images*}
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   572
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   573
lemma image_pair_mono:
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   574
    "[| !! x y. <x,y>:r ==> <x,y>:s;  A<=B |] ==> r``A <= s``B"
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   575
by blast 
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   576
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   577
lemma vimage_pair_mono:
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   578
    "[| !! x y. <x,y>:r ==> <x,y>:s;  A<=B |] ==> r-``A <= s-``B"
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   579
by blast 
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   580
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   581
lemma image_mono: "[| r<=s;  A<=B |] ==> r``A <= s``B"
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 vimage_mono: "[| r<=s;  A<=B |] ==> r-``A <= s-``B"
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   585
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   586
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   587
lemma Collect_mono:
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   588
    "[| A<=B;  !!x. x:A ==> P(x) --> Q(x) |] ==> Collect(A,P) <= Collect(B,Q)"
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   589
by blast
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   590
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   591
(*Used in intr_elim.ML and in individual datatype definitions*)
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   592
lemmas basic_monos = subset_refl imp_refl disj_mono conj_mono ex_mono 
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   593
                     Collect_mono Part_mono in_mono
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   594
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   595
ML
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   596
{*
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   597
val Pi_iff = thm "Pi_iff";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   598
val Pi_iff_old = thm "Pi_iff_old";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   599
val fun_is_function = thm "fun_is_function";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   600
val fun_is_rel = thm "fun_is_rel";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   601
val Pi_cong = thm "Pi_cong";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   602
val fun_weaken_type = thm "fun_weaken_type";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   603
val apply_equality2 = thm "apply_equality2";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   604
val function_apply_equality = thm "function_apply_equality";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   605
val apply_equality = thm "apply_equality";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   606
val apply_0 = thm "apply_0";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   607
val Pi_memberD = thm "Pi_memberD";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   608
val function_apply_Pair = thm "function_apply_Pair";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   609
val apply_Pair = thm "apply_Pair";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   610
val apply_type = thm "apply_type";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   611
val apply_funtype = thm "apply_funtype";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   612
val apply_iff = thm "apply_iff";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   613
val Pi_type = thm "Pi_type";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   614
val Pi_Collect_iff = thm "Pi_Collect_iff";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   615
val Pi_weaken_type = thm "Pi_weaken_type";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   616
val domain_type = thm "domain_type";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   617
val range_type = thm "range_type";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   618
val Pair_mem_PiD = thm "Pair_mem_PiD";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   619
val lamI = thm "lamI";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   620
val lamE = thm "lamE";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   621
val lamD = thm "lamD";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   622
val lam_type = thm "lam_type";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   623
val lam_funtype = thm "lam_funtype";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   624
val beta = thm "beta";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   625
val lam_empty = thm "lam_empty";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   626
val domain_lam = thm "domain_lam";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   627
val lam_cong = thm "lam_cong";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   628
val lam_theI = thm "lam_theI";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   629
val lam_eqE = thm "lam_eqE";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   630
val Pi_empty1 = thm "Pi_empty1";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   631
val singleton_fun = thm "singleton_fun";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   632
val Pi_empty2 = thm "Pi_empty2";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   633
val fun_space_empty_iff = thm "fun_space_empty_iff";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   634
val fun_subset = thm "fun_subset";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   635
val fun_extension = thm "fun_extension";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   636
val eta = thm "eta";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   637
val fun_extension_iff = thm "fun_extension_iff";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   638
val fun_subset_eq = thm "fun_subset_eq";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   639
val Pi_lamE = thm "Pi_lamE";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   640
val image_lam = thm "image_lam";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   641
val image_fun = thm "image_fun";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   642
val Pi_image_cons = thm "Pi_image_cons";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   643
val restrict_subset = thm "restrict_subset";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   644
val function_restrictI = thm "function_restrictI";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   645
val restrict_type2 = thm "restrict_type2";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   646
val restrict = thm "restrict";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   647
val restrict_empty = thm "restrict_empty";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   648
val domain_restrict_lam = thm "domain_restrict_lam";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   649
val restrict_restrict = thm "restrict_restrict";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   650
val domain_restrict = thm "domain_restrict";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   651
val restrict_idem = thm "restrict_idem";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   652
val restrict_if = thm "restrict_if";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   653
val restrict_lam_eq = thm "restrict_lam_eq";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   654
val fun_cons_restrict_eq = thm "fun_cons_restrict_eq";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   655
val function_Union = thm "function_Union";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   656
val fun_Union = thm "fun_Union";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   657
val fun_disjoint_Un = thm "fun_disjoint_Un";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   658
val fun_disjoint_apply1 = thm "fun_disjoint_apply1";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   659
val fun_disjoint_apply2 = thm "fun_disjoint_apply2";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   660
val domain_of_fun = thm "domain_of_fun";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   661
val apply_rangeI = thm "apply_rangeI";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   662
val range_of_fun = thm "range_of_fun";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   663
val fun_extend = thm "fun_extend";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   664
val fun_extend3 = thm "fun_extend3";
13176
312bd350579b conversion of Perm to Isar. Strengthening of comp_fun_apply
paulson
parents: 13175
diff changeset
   665
val fun_extend_apply = thm "fun_extend_apply";
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   666
val singleton_apply = thm "singleton_apply";
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   667
val cons_fun_eq = thm "cons_fun_eq";
13355
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   668
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   669
val update_def = thm "update_def";
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   670
val update_apply = thm "update_apply";
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   671
val update_idem = thm "update_idem";
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   672
val domain_update = thm "domain_update";
d19cdbd8b559 merged Update with func
paulson
parents: 13269
diff changeset
   673
val update_type = thm "update_type";
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   674
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   675
val Replace_mono = thm "Replace_mono";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   676
val RepFun_mono = thm "RepFun_mono";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   677
val Pow_mono = thm "Pow_mono";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   678
val Union_mono = thm "Union_mono";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   679
val UN_mono = thm "UN_mono";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   680
val Inter_anti_mono = thm "Inter_anti_mono";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   681
val cons_mono = thm "cons_mono";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   682
val Un_mono = thm "Un_mono";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   683
val Int_mono = thm "Int_mono";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   684
val Diff_mono = thm "Diff_mono";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   685
val Sigma_mono = thm "Sigma_mono";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   686
val sum_mono = thm "sum_mono";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   687
val Pi_mono = thm "Pi_mono";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   688
val lam_mono = thm "lam_mono";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   689
val converse_mono = thm "converse_mono";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   690
val domain_mono = thm "domain_mono";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   691
val domain_rel_subset = thm "domain_rel_subset";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   692
val range_mono = thm "range_mono";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   693
val range_rel_subset = thm "range_rel_subset";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   694
val field_mono = thm "field_mono";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   695
val field_rel_subset = thm "field_rel_subset";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   696
val image_pair_mono = thm "image_pair_mono";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   697
val vimage_pair_mono = thm "vimage_pair_mono";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   698
val image_mono = thm "image_mono";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   699
val vimage_mono = thm "vimage_mono";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   700
val Collect_mono = thm "Collect_mono";
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   701
6f54e992777e Removal of mono.thy
paulson
parents: 13355
diff changeset
   702
val basic_monos = thms "basic_monos";
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   703
*}
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   704
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 2469
diff changeset
   705
end