src/HOL/Analysis/Cross3.thy
author wenzelm
Wed, 02 Oct 2024 19:55:07 +0200
changeset 81104 56eebde7ac7e
parent 81100 6ae3d0b2b8ad
child 81111 f1a3a553e8cf
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
68466
3d8241f4198b corrections to markup
paulson <lp15@cam.ac.uk>
parents: 68465
diff changeset
     1
(* Title:      HOL/Analysis/Cross3.thy
3d8241f4198b corrections to markup
paulson <lp15@cam.ac.uk>
parents: 68465
diff changeset
     2
   Author:     L C Paulson, University of Cambridge
3d8241f4198b corrections to markup
paulson <lp15@cam.ac.uk>
parents: 68465
diff changeset
     3
3d8241f4198b corrections to markup
paulson <lp15@cam.ac.uk>
parents: 68465
diff changeset
     4
Ported from HOL Light
3d8241f4198b corrections to markup
paulson <lp15@cam.ac.uk>
parents: 68465
diff changeset
     5
*)
3d8241f4198b corrections to markup
paulson <lp15@cam.ac.uk>
parents: 68465
diff changeset
     6
68902
nipkow
parents: 68833
diff changeset
     7
section\<open>Vector Cross Products in 3 Dimensions\<close>
68466
3d8241f4198b corrections to markup
paulson <lp15@cam.ac.uk>
parents: 68465
diff changeset
     8
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     9
theory "Cross3"
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69508
diff changeset
    10
  imports Determinants Cartesian_Euclidean_Space
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    11
begin
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    12
68467
44ffc5b9cd76 fixing overloading problems involving vector cross products
paulson <lp15@cam.ac.uk>
parents: 68466
diff changeset
    13
context includes no_Set_Product_syntax 
44ffc5b9cd76 fixing overloading problems involving vector cross products
paulson <lp15@cam.ac.uk>
parents: 68466
diff changeset
    14
begin \<comment>\<open>locally disable syntax for set product, to avoid warnings\<close>
44ffc5b9cd76 fixing overloading problems involving vector cross products
paulson <lp15@cam.ac.uk>
parents: 68466
diff changeset
    15
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 73932
diff changeset
    16
definition\<^marker>\<open>tag important\<close> cross3 :: "[real^3, real^3] \<Rightarrow> real^3"  (infixr \<open>\<times>\<close> 80)
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    17
  where "a \<times> b \<equiv>
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    18
    vector [a$2 * b$3 - a$3 * b$2,
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    19
            a$3 * b$1 - a$1 * b$3,
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    20
            a$1 * b$2 - a$2 * b$1]"
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    21
68467
44ffc5b9cd76 fixing overloading problems involving vector cross products
paulson <lp15@cam.ac.uk>
parents: 68466
diff changeset
    22
end
44ffc5b9cd76 fixing overloading problems involving vector cross products
paulson <lp15@cam.ac.uk>
parents: 68466
diff changeset
    23
81100
6ae3d0b2b8ad tuned whitespace;
wenzelm
parents: 80914
diff changeset
    24
bundle cross3_syntax
6ae3d0b2b8ad tuned whitespace;
wenzelm
parents: 80914
diff changeset
    25
begin
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 73932
diff changeset
    26
notation cross3 (infixr \<open>\<times>\<close> 80)
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 73932
diff changeset
    27
no_notation Product_Type.Times (infixr \<open>\<times>\<close> 80)
68467
44ffc5b9cd76 fixing overloading problems involving vector cross products
paulson <lp15@cam.ac.uk>
parents: 68466
diff changeset
    28
end
44ffc5b9cd76 fixing overloading problems involving vector cross products
paulson <lp15@cam.ac.uk>
parents: 68466
diff changeset
    29
81100
6ae3d0b2b8ad tuned whitespace;
wenzelm
parents: 80914
diff changeset
    30
bundle no_cross3_syntax
6ae3d0b2b8ad tuned whitespace;
wenzelm
parents: 80914
diff changeset
    31
begin
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 73932
diff changeset
    32
no_notation cross3 (infixr \<open>\<times>\<close> 80)
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 73932
diff changeset
    33
notation Product_Type.Times (infixr \<open>\<times>\<close> 80)
68467
44ffc5b9cd76 fixing overloading problems involving vector cross products
paulson <lp15@cam.ac.uk>
parents: 68466
diff changeset
    34
end
44ffc5b9cd76 fixing overloading problems involving vector cross products
paulson <lp15@cam.ac.uk>
parents: 68466
diff changeset
    35
44ffc5b9cd76 fixing overloading problems involving vector cross products
paulson <lp15@cam.ac.uk>
parents: 68466
diff changeset
    36
unbundle cross3_syntax
44ffc5b9cd76 fixing overloading problems involving vector cross products
paulson <lp15@cam.ac.uk>
parents: 68466
diff changeset
    37
69683
8b3458ca0762 subsection is always %important
immler
parents: 69682
diff changeset
    38
subsection\<open> Basic lemmas\<close>
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    39
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    40
lemmas cross3_simps = cross3_def inner_vec_def sum_3 det_3 vec_eq_iff vector_def algebra_simps
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    41
68902
nipkow
parents: 68833
diff changeset
    42
lemma dot_cross_self: "x \<bullet> (x \<times> y) = 0" "x \<bullet> (y \<times> x) = 0" "(x \<times> y) \<bullet> y = 0" "(y \<times> x) \<bullet> y = 0"
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    43
  by (simp_all add: orthogonal_def cross3_simps)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    44
68902
nipkow
parents: 68833
diff changeset
    45
lemma  orthogonal_cross: "orthogonal (x \<times> y) x" "orthogonal (x \<times> y) y"  
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    46
                        "orthogonal y (x \<times> y)" "orthogonal (x \<times> y) x"
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    47
  by (simp_all add: orthogonal_def dot_cross_self)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    48
68902
nipkow
parents: 68833
diff changeset
    49
lemma  cross_zero_left [simp]: "0 \<times> x = 0" and cross_zero_right [simp]: "x \<times> 0 = 0" for x::"real^3"
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    50
  by (simp_all add: cross3_simps)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    51
68902
nipkow
parents: 68833
diff changeset
    52
lemma  cross_skew: "(x \<times> y) = -(y \<times> x)" for x::"real^3"
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    53
  by (simp add: cross3_simps)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    54
68902
nipkow
parents: 68833
diff changeset
    55
lemma  cross_refl [simp]: "x \<times> x = 0" for x::"real^3"
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    56
  by (simp add: cross3_simps)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
68902
nipkow
parents: 68833
diff changeset
    58
lemma  cross_add_left: "(x + y) \<times> z = (x \<times> z) + (y \<times> z)" for x::"real^3"
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    59
  by (simp add: cross3_simps)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    60
68902
nipkow
parents: 68833
diff changeset
    61
lemma  cross_add_right: "x \<times> (y + z) = (x \<times> y) + (x \<times> z)" for x::"real^3"
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    62
  by (simp add: cross3_simps)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    63
68902
nipkow
parents: 68833
diff changeset
    64
lemma  cross_mult_left: "(c *\<^sub>R x) \<times> y = c *\<^sub>R (x \<times> y)" for x::"real^3"
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    65
  by (simp add: cross3_simps)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    66
68902
nipkow
parents: 68833
diff changeset
    67
lemma  cross_mult_right: "x \<times> (c *\<^sub>R y) = c *\<^sub>R (x \<times> y)" for x::"real^3"
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
  by (simp add: cross3_simps)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    69
68902
nipkow
parents: 68833
diff changeset
    70
lemma  cross_minus_left [simp]: "(-x) \<times> y = - (x \<times> y)" for x::"real^3"
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    71
  by (simp add: cross3_simps)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    72
68902
nipkow
parents: 68833
diff changeset
    73
lemma  cross_minus_right [simp]: "x \<times> -y = - (x \<times> y)" for x::"real^3"
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    74
  by (simp add: cross3_simps)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    75
68902
nipkow
parents: 68833
diff changeset
    76
lemma  left_diff_distrib: "(x - y) \<times> z = x \<times> z - y \<times> z" for x::"real^3"
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
  by (simp add: cross3_simps)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    78
68902
nipkow
parents: 68833
diff changeset
    79
lemma  right_diff_distrib: "x \<times> (y - z) = x \<times> y - x \<times> z" for x::"real^3"
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
  by (simp add: cross3_simps)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    81
68467
44ffc5b9cd76 fixing overloading problems involving vector cross products
paulson <lp15@cam.ac.uk>
parents: 68466
diff changeset
    82
hide_fact (open) left_diff_distrib right_diff_distrib
44ffc5b9cd76 fixing overloading problems involving vector cross products
paulson <lp15@cam.ac.uk>
parents: 68466
diff changeset
    83
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
    84
proposition Jacobi: "x \<times> (y \<times> z) + y \<times> (z \<times> x) + z \<times> (x \<times> y) = 0" for x::"real^3"
69682
500a7acdccd4 no need for %unimportant for proofs of proposition
immler
parents: 69681
diff changeset
    85
  by (simp add: cross3_simps)
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    86
69682
500a7acdccd4 no need for %unimportant for proofs of proposition
immler
parents: 69681
diff changeset
    87
proposition Lagrange: "x \<times> (y \<times> z) = (x \<bullet> z) *\<^sub>R y - (x \<bullet> y) *\<^sub>R z"
500a7acdccd4 no need for %unimportant for proofs of proposition
immler
parents: 69681
diff changeset
    88
  by (simp add: cross3_simps) (metis (full_types) exhaust_3)
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    89
69682
500a7acdccd4 no need for %unimportant for proofs of proposition
immler
parents: 69681
diff changeset
    90
proposition cross_triple: "(x \<times> y) \<bullet> z = (y \<times> z) \<bullet> x"
500a7acdccd4 no need for %unimportant for proofs of proposition
immler
parents: 69681
diff changeset
    91
  by (simp add: cross3_def inner_vec_def sum_3 vec_eq_iff algebra_simps)
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
68902
nipkow
parents: 68833
diff changeset
    93
lemma  cross_components:
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    94
   "(x \<times> y)$1 = x$2 * y$3 - y$2 * x$3" "(x \<times> y)$2 = x$3 * y$1 - y$3 * x$1" "(x \<times> y)$3 = x$1 * y$2 - y$1 * x$2"
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    95
  by (simp_all add: cross3_def inner_vec_def sum_3 vec_eq_iff algebra_simps)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    96
68902
nipkow
parents: 68833
diff changeset
    97
lemma  cross_basis: "(axis 1 1) \<times> (axis 2 1) = axis 3 1" "(axis 2 1) \<times> (axis 1 1) = -(axis 3 1)" 
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    98
                   "(axis 2 1) \<times> (axis 3 1) = axis 1 1" "(axis 3 1) \<times> (axis 2 1) = -(axis 1 1)" 
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
                   "(axis 3 1) \<times> (axis 1 1) = axis 2 1" "(axis 1 1) \<times> (axis 3 1) = -(axis 2 1)"
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   100
  using exhaust_3
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   101
  by (force simp add: axis_def cross3_simps)+
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   102
68902
nipkow
parents: 68833
diff changeset
   103
lemma  cross_basis_nonzero:
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 68902
diff changeset
   104
  "u \<noteq> 0 \<Longrightarrow> u \<times> axis 1 1 \<noteq> 0 \<or> u \<times> axis 2 1 \<noteq> 0 \<or> u \<times> axis 3 1 \<noteq> 0"
71633
07bec530f02e cleaned proofs
nipkow
parents: 70136
diff changeset
   105
  by (clarsimp simp add: axis_def cross3_simps) (metis exhaust_3)
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   106
68902
nipkow
parents: 68833
diff changeset
   107
lemma  cross_dot_cancel:
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   108
  fixes x::"real^3"
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   109
  assumes deq: "x \<bullet> y = x \<bullet> z" and veq: "x \<times> y = x \<times> z" and x: "x \<noteq> 0"
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   110
  shows "y = z" 
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   111
proof -
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   112
  have "x \<bullet> x \<noteq> 0"
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   113
    by (simp add: x)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   114
  then have "y - z = 0"
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   115
    using veq
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   116
    by (metis (no_types, lifting) Cross3.right_diff_distrib Lagrange deq eq_iff_diff_eq_0 inner_diff_right scale_eq_0_iff)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   117
  then show ?thesis
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   118
    using eq_iff_diff_eq_0 by blast
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   119
qed
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   120
68902
nipkow
parents: 68833
diff changeset
   121
lemma  norm_cross_dot: "(norm (x \<times> y))\<^sup>2 + (x \<bullet> y)\<^sup>2 = (norm x * norm y)\<^sup>2"
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   122
  unfolding power2_norm_eq_inner power_mult_distrib
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   123
  by (simp add: cross3_simps power2_eq_square)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   124
68902
nipkow
parents: 68833
diff changeset
   125
lemma  dot_cross_det: "x \<bullet> (y \<times> z) = det(vector[x,y,z])"
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
  by (simp add: cross3_simps) 
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   127
68902
nipkow
parents: 68833
diff changeset
   128
lemma  cross_cross_det: "(w \<times> x) \<times> (y \<times> z) = det(vector[w,x,z]) *\<^sub>R y - det(vector[w,x,y]) *\<^sub>R z"
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   129
  using exhaust_3 by (force simp add: cross3_simps) 
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   130
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   131
proposition  dot_cross: "(w \<times> x) \<bullet> (y \<times> z) = (w \<bullet> y) * (x \<bullet> z) - (w \<bullet> z) * (x \<bullet> y)"
69682
500a7acdccd4 no need for %unimportant for proofs of proposition
immler
parents: 69681
diff changeset
   132
  by (force simp add: cross3_simps)
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   133
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   134
proposition  norm_cross: "(norm (x \<times> y))\<^sup>2 = (norm x)\<^sup>2 * (norm y)\<^sup>2 - (x \<bullet> y)\<^sup>2"
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   135
  unfolding power2_norm_eq_inner power_mult_distrib
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   136
  by (simp add: cross3_simps power2_eq_square)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   137
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   138
lemma  cross_eq_0: "x \<times> y = 0 \<longleftrightarrow> collinear{0,x,y}"
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   139
proof -
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   140
  have "x \<times> y = 0 \<longleftrightarrow> norm (x \<times> y) = 0"
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   141
    by simp
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   142
  also have "... \<longleftrightarrow> (norm x * norm y)\<^sup>2 = (x \<bullet> y)\<^sup>2"
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   143
    using norm_cross [of x y] by (auto simp: power_mult_distrib)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   144
  also have "... \<longleftrightarrow> \<bar>x \<bullet> y\<bar> = norm x * norm y"
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   145
    using power2_eq_iff
73932
fd21b4a93043 added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents: 71633
diff changeset
   146
    by (metis (mono_tags, opaque_lifting) abs_minus abs_norm_cancel abs_power2 norm_mult power_abs real_norm_def) 
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   147
  also have "... \<longleftrightarrow> collinear {0, x, y}"
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   148
    by (rule norm_cauchy_schwarz_equal)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   149
  finally show ?thesis .
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   150
qed
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   151
68902
nipkow
parents: 68833
diff changeset
   152
lemma  cross_eq_self: "x \<times> y = x \<longleftrightarrow> x = 0" "x \<times> y = y \<longleftrightarrow> y = 0"
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   153
  apply (metis cross_zero_left dot_cross_self(1) inner_eq_zero_iff)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   154
  by (metis cross_zero_right dot_cross_self(2) inner_eq_zero_iff)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   155
68902
nipkow
parents: 68833
diff changeset
   156
lemma  norm_and_cross_eq_0:
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   157
   "x \<bullet> y = 0 \<and> x \<times> y = 0 \<longleftrightarrow> x = 0 \<or> y = 0" (is "?lhs = ?rhs")
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   158
proof 
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   159
  assume ?lhs
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   160
  then show ?rhs
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   161
    by (metis cross_dot_cancel cross_zero_right inner_zero_right)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   162
qed auto
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   163
68902
nipkow
parents: 68833
diff changeset
   164
lemma  bilinear_cross: "bilinear(\<times>)"
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   165
  apply (auto simp add: bilinear_def linear_def)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   166
  apply unfold_locales
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   167
  apply (simp add: cross_add_right)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   168
  apply (simp add: cross_mult_right)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   169
  apply (simp add: cross_add_left)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   170
  apply (simp add: cross_mult_left)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   171
  done
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   172
69683
8b3458ca0762 subsection is always %important
immler
parents: 69682
diff changeset
   173
subsection   \<open>Preservation by rotation, or other orthogonal transformation up to sign\<close>
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   174
68902
nipkow
parents: 68833
diff changeset
   175
lemma  cross_matrix_mult: "transpose A *v ((A *v x) \<times> (A *v y)) = det A *\<^sub>R (x \<times> y)"
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   176
  apply (simp add: vec_eq_iff   )
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   177
  apply (simp add: vector_matrix_mult_def matrix_vector_mult_def forall_3 cross3_simps)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   178
  done
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   179
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   180
lemma  cross_orthogonal_matrix:
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   181
  assumes "orthogonal_matrix A"
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   182
  shows "(A *v x) \<times> (A *v y) = det A *\<^sub>R (A *v (x \<times> y))"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   183
proof -
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   184
  have "mat 1 = transpose (A ** transpose A)"
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   185
    by (metis (no_types) assms orthogonal_matrix_def transpose_mat)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   186
  then show ?thesis
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   187
    by (metis (no_types) vector_matrix_mul_rid vector_transpose_matrix cross_matrix_mult matrix_vector_mul_assoc matrix_vector_mult_scaleR)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   188
qed
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   189
68902
nipkow
parents: 68833
diff changeset
   190
lemma  cross_rotation_matrix: "rotation_matrix A \<Longrightarrow> (A *v x) \<times> (A *v y) =  A *v (x \<times> y)"
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   191
  by (simp add: rotation_matrix_def cross_orthogonal_matrix)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   192
68902
nipkow
parents: 68833
diff changeset
   193
lemma  cross_rotoinversion_matrix: "rotoinversion_matrix A \<Longrightarrow> (A *v x) \<times> (A *v y) = - A *v (x \<times> y)"
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   194
  by (simp add: rotoinversion_matrix_def cross_orthogonal_matrix scaleR_matrix_vector_assoc)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   195
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   196
lemma  cross_orthogonal_transformation:
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   197
  assumes "orthogonal_transformation f"
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   198
  shows   "(f x) \<times> (f y) = det(matrix f) *\<^sub>R f(x \<times> y)"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   199
proof -
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   200
  have orth: "orthogonal_matrix (matrix f)"
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   201
    using assms orthogonal_transformation_matrix by blast
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   202
  have "matrix f *v z = f z" for z
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   203
    using assms orthogonal_transformation_matrix by force
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   204
  with cross_orthogonal_matrix [OF orth] show ?thesis
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   205
    by simp
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   206
qed
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   207
68902
nipkow
parents: 68833
diff changeset
   208
lemma  cross_linear_image:
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   209
   "\<lbrakk>linear f; \<And>x. norm(f x) = norm x; det(matrix f) = 1\<rbrakk>
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   210
           \<Longrightarrow> (f x) \<times> (f y) = f(x \<times> y)"
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   211
  by (simp add: cross_orthogonal_transformation orthogonal_transformation)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   212
69683
8b3458ca0762 subsection is always %important
immler
parents: 69682
diff changeset
   213
subsection \<open>Continuity\<close>
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   214
68902
nipkow
parents: 68833
diff changeset
   215
lemma  continuous_cross: "\<lbrakk>continuous F f; continuous F g\<rbrakk> \<Longrightarrow> continuous F (\<lambda>x. (f x) \<times> (g x))"
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   216
  apply (subst continuous_componentwise)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   217
  apply (clarsimp simp add: cross3_simps)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   218
  apply (intro continuous_intros; simp)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   219
  done
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   220
68902
nipkow
parents: 68833
diff changeset
   221
lemma  continuous_on_cross:
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   222
  fixes f :: "'a::t2_space \<Rightarrow> real^3"
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   223
  shows "\<lbrakk>continuous_on S f; continuous_on S g\<rbrakk> \<Longrightarrow> continuous_on S (\<lambda>x. (f x) \<times> (g x))"
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   224
  by (simp add: continuous_on_eq_continuous_within continuous_cross)
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   225
68467
44ffc5b9cd76 fixing overloading problems involving vector cross products
paulson <lp15@cam.ac.uk>
parents: 68466
diff changeset
   226
unbundle no_cross3_syntax
44ffc5b9cd76 fixing overloading problems involving vector cross products
paulson <lp15@cam.ac.uk>
parents: 68466
diff changeset
   227
68465
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   228
end
e699ca8e22b7 New material in support of quaternions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   229