src/HOL/Library/Function_Division.thy
author wenzelm
Mon, 14 Mar 2016 17:43:17 +0100
changeset 62628 6031191a8d9c
parent 60500 903bb1495239
child 69593 3dda49e08b9d
permissions -rw-r--r--
record stamps of cumulative input heaps; tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
48188
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/Library/Function_Division.thy
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
     2
    Author:     Florian Haftmann, TUM
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
     3
*)
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
     4
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60429
diff changeset
     5
section \<open>Pointwise instantiation of functions to division\<close>
48188
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
     6
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
     7
theory Function_Division
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
     8
imports Function_Algebras
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
     9
begin
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    10
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60429
diff changeset
    11
subsection \<open>Syntactic with division\<close>
48188
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    12
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    13
instantiation "fun" :: (type, inverse) inverse
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    14
begin
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    15
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    16
definition "inverse f = inverse \<circ> f"
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    17
60429
d3d1e185cd63 uniform _ div _ as infix syntax for ring division
haftmann
parents: 60352
diff changeset
    18
definition "f div g = (\<lambda>x. f x / g x)"
48188
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    19
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    20
instance ..
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    21
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    22
end
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    23
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    24
lemma inverse_fun_apply [simp]:
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    25
  "inverse f x = inverse (f x)"
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    26
  by (simp add: inverse_fun_def)
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    27
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    28
lemma divide_fun_apply [simp]:
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    29
  "(f / g) x = f x / g x"
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    30
  by (simp add: divide_fun_def)
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    31
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60429
diff changeset
    32
text \<open>
48188
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    33
  Unfortunately, we cannot lift this operations to algebraic type
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    34
  classes for division: being different from the constant
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    35
  zero function @{term "f \<noteq> 0"} is too weak as precondition.
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    36
  So we must introduce our own set of lemmas.
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60429
diff changeset
    37
\<close>
48188
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    38
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    39
abbreviation zero_free :: "('b \<Rightarrow> 'a::field) \<Rightarrow> bool" where
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    40
  "zero_free f \<equiv> \<not> (\<exists>x. f x = 0)"
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    41
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    42
lemma fun_left_inverse:
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    43
  fixes f :: "'b \<Rightarrow> 'a::field"
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    44
  shows "zero_free f \<Longrightarrow> inverse f * f = 1"
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    45
  by (simp add: fun_eq_iff)
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    46
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    47
lemma fun_right_inverse:
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    48
  fixes f :: "'b \<Rightarrow> 'a::field"
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    49
  shows "zero_free f \<Longrightarrow> f * inverse f = 1"
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    50
  by (simp add: fun_eq_iff)
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    51
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    52
lemma fun_divide_inverse:
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    53
  fixes f g :: "'b \<Rightarrow> 'a::field"
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    54
  shows "f / g = f * inverse g"
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    55
  by (simp add: fun_eq_iff divide_inverse)
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    56
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60429
diff changeset
    57
text \<open>Feel free to extend this.\<close>
48188
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    58
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60429
diff changeset
    59
text \<open>
48188
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    60
  Another possibility would be a reformulation of the division type
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    61
  classes to user a @{term zero_free} predicate rather than
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    62
  a direct @{term "a \<noteq> 0"} condition.
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60429
diff changeset
    63
\<close>
48188
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    64
dcfe2c92fc7c Stub theory for division on functionals.
haftmann
parents:
diff changeset
    65
end