src/HOL/Decision_Procs/Reflective_Field.thy
author wenzelm
Sun, 24 Oct 2021 20:25:51 +0200
changeset 74570 7625b5d7cfe2
parent 74569 f4613ca298e6
child 78095 bc42c074e58f
permissions -rw-r--r--
clarified name, syntax, messages;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
     1
(*  Title:      HOL/Decision_Procs/Reflective_Field.thy
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
     2
    Author:     Stefan Berghofer
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
     3
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
     4
Reducing equalities in fields to equalities in rings.
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
     5
*)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
     6
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
     7
theory Reflective_Field
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
     8
  imports Commutative_Ring
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
     9
begin
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    10
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    11
datatype fexpr =
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    12
    FCnst int
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    13
  | FVar nat
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    14
  | FAdd fexpr fexpr
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    15
  | FSub fexpr fexpr
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    16
  | FMul fexpr fexpr
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    17
  | FNeg fexpr
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    18
  | FDiv fexpr fexpr
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    19
  | FPow fexpr nat
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    20
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    21
fun (in field) nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    22
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    23
    "nth_el [] n = \<zero>"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    24
  | "nth_el (x # xs) 0 = x"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    25
  | "nth_el (x # xs) (Suc n) = nth_el xs n"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    26
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    27
lemma (in field) nth_el_Cons: "nth_el (x # xs) n = (if n = 0 then x else nth_el xs (n - 1))"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    28
  by (cases n) simp_all
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    29
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    30
lemma (in field) nth_el_closed [simp]: "in_carrier xs \<Longrightarrow> nth_el xs n \<in> carrier R"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    31
  by (induct xs n rule: nth_el.induct) (simp_all add: in_carrier_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    32
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    33
primrec (in field) feval :: "'a list \<Rightarrow> fexpr \<Rightarrow> 'a"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    34
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    35
    "feval xs (FCnst c) = \<guillemotleft>c\<guillemotright>"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    36
  | "feval xs (FVar n) = nth_el xs n"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    37
  | "feval xs (FAdd a b) = feval xs a \<oplus> feval xs b"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    38
  | "feval xs (FSub a b) = feval xs a \<ominus> feval xs b"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    39
  | "feval xs (FMul a b) = feval xs a \<otimes> feval xs b"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    40
  | "feval xs (FNeg a) = \<ominus> feval xs a"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    41
  | "feval xs (FDiv a b) = feval xs a \<oslash> feval xs b"
67341
df79ef3b3a41 Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents: 67123
diff changeset
    42
  | "feval xs (FPow a n) = feval xs a [^] n"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    43
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    44
lemma (in field) feval_Cnst:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    45
  "feval xs (FCnst 0) = \<zero>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    46
  "feval xs (FCnst 1) = \<one>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    47
  "feval xs (FCnst (numeral n)) = \<guillemotleft>numeral n\<guillemotright>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    48
  by simp_all
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    49
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    50
datatype pexpr =
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    51
    PExpr1 pexpr1
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    52
  | PExpr2 pexpr2
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    53
and pexpr1 =
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    54
    PCnst int
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    55
  | PVar nat
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    56
  | PAdd pexpr pexpr
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    57
  | PSub pexpr pexpr
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    58
  | PNeg pexpr
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    59
and pexpr2 =
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    60
    PMul pexpr pexpr
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    61
  | PPow pexpr nat
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    62
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    63
lemma pexpr_cases [case_names PCnst PVar PAdd PSub PNeg PMul PPow]:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    64
  assumes
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    65
    "\<And>c. e = PExpr1 (PCnst c) \<Longrightarrow> P"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    66
    "\<And>n. e = PExpr1 (PVar n) \<Longrightarrow> P"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    67
    "\<And>e1 e2. e = PExpr1 (PAdd e1 e2) \<Longrightarrow> P"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    68
    "\<And>e1 e2. e = PExpr1 (PSub e1 e2) \<Longrightarrow> P"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    69
    "\<And>e'. e = PExpr1 (PNeg e') \<Longrightarrow> P"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    70
    "\<And>e1 e2. e = PExpr2 (PMul e1 e2) \<Longrightarrow> P"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    71
    "\<And>e' n. e = PExpr2 (PPow e' n) \<Longrightarrow> P"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    72
  shows P
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    73
proof (cases e)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    74
  case (PExpr1 e')
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    75
  then show ?thesis
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    76
    apply (cases e')
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    77
        apply simp_all
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    78
        apply (erule assms)+
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    79
    done
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    80
next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    81
  case (PExpr2 e')
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    82
  then show ?thesis
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    83
    apply (cases e')
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    84
     apply simp_all
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    85
     apply (erule assms)+
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    86
    done
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    87
qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    88
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    89
lemmas pexpr_cases2 = pexpr_cases [case_product pexpr_cases]
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    90
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    91
fun (in field) peval :: "'a list \<Rightarrow> pexpr \<Rightarrow> 'a"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    92
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    93
    "peval xs (PExpr1 (PCnst c)) = \<guillemotleft>c\<guillemotright>"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    94
  | "peval xs (PExpr1 (PVar n)) = nth_el xs n"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    95
  | "peval xs (PExpr1 (PAdd a b)) = peval xs a \<oplus> peval xs b"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    96
  | "peval xs (PExpr1 (PSub a b)) = peval xs a \<ominus> peval xs b"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    97
  | "peval xs (PExpr1 (PNeg a)) = \<ominus> peval xs a"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
    98
  | "peval xs (PExpr2 (PMul a b)) = peval xs a \<otimes> peval xs b"
67341
df79ef3b3a41 Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents: 67123
diff changeset
    99
  | "peval xs (PExpr2 (PPow a n)) = peval xs a [^] n"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   100
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   101
lemma (in field) peval_Cnst:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   102
  "peval xs (PExpr1 (PCnst 0)) = \<zero>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   103
  "peval xs (PExpr1 (PCnst 1)) = \<one>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   104
  "peval xs (PExpr1 (PCnst (numeral n))) = \<guillemotleft>numeral n\<guillemotright>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   105
  "peval xs (PExpr1 (PCnst (- numeral n))) = \<ominus> \<guillemotleft>numeral n\<guillemotright>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   106
  by simp_all
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   107
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   108
lemma (in field) peval_closed [simp]:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   109
  "in_carrier xs \<Longrightarrow> peval xs e \<in> carrier R"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   110
  "in_carrier xs \<Longrightarrow> peval xs (PExpr1 e1) \<in> carrier R"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   111
  "in_carrier xs \<Longrightarrow> peval xs (PExpr2 e2) \<in> carrier R"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   112
  by (induct e and e1 and e2) simp_all
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   113
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   114
definition npepow :: "pexpr \<Rightarrow> nat \<Rightarrow> pexpr"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   115
  where "npepow e n =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   116
   (if n = 0 then PExpr1 (PCnst 1)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   117
    else if n = 1 then e
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   118
    else
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   119
      (case e of
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   120
        PExpr1 (PCnst c) \<Rightarrow> PExpr1 (PCnst (c ^ n))
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   121
      | _ \<Rightarrow> PExpr2 (PPow e n)))"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   122
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   123
lemma (in field) npepow_correct:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   124
  "in_carrier xs \<Longrightarrow> peval xs (npepow e n) = peval xs (PExpr2 (PPow e n))"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   125
  by (cases e rule: pexpr_cases) (simp_all add: npepow_def)
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   126
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   127
fun npemul :: "pexpr \<Rightarrow> pexpr \<Rightarrow> pexpr"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   128
  where "npemul x y =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   129
    (case x of
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   130
      PExpr1 (PCnst c) \<Rightarrow>
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   131
        if c = 0 then x
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   132
        else if c = 1 then y else
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   133
          (case y of
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   134
            PExpr1 (PCnst d) \<Rightarrow> PExpr1 (PCnst (c * d))
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   135
          | _ \<Rightarrow> PExpr2 (PMul x y))
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   136
    | PExpr2 (PPow e1 n) \<Rightarrow>
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   137
        (case y of
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   138
          PExpr2 (PPow e2 m) \<Rightarrow>
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   139
            if n = m then npepow (npemul e1 e2) n
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   140
            else PExpr2 (PMul x y)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   141
        | PExpr1 (PCnst d) \<Rightarrow>
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   142
            if d = 0 then y
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   143
            else if d = 1 then x
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   144
            else PExpr2 (PMul x y)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   145
        | _ \<Rightarrow> PExpr2 (PMul x y))
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   146
    | _ \<Rightarrow>
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   147
      (case y of
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   148
        PExpr1 (PCnst d) \<Rightarrow>
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   149
          if d = 0 then y
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   150
          else if d = 1 then x
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   151
          else PExpr2 (PMul x y)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   152
      | _ \<Rightarrow> PExpr2 (PMul x y)))"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   153
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   154
lemma (in field) npemul_correct:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   155
  "in_carrier xs \<Longrightarrow> peval xs (npemul e1 e2) = peval xs (PExpr2 (PMul e1 e2))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   156
proof (induct e1 e2 rule: npemul.induct)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   157
  case (1 x y)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   158
  then show ?case
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   159
  proof (cases x y rule: pexpr_cases2)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   160
    case (PPow_PPow e n e' m)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   161
    then show ?thesis
70019
095dce9892e8 A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   162
      by (simp del: npemul.simps add: 1 npepow_correct nat_pow_distrib
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   163
          npemul.simps [of "PExpr2 (PPow e n)" "PExpr2 (PPow e' m)"])
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   164
  qed simp_all
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   165
qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   166
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   167
declare npemul.simps [simp del]
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   168
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   169
definition npeadd :: "pexpr \<Rightarrow> pexpr \<Rightarrow> pexpr"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   170
  where "npeadd x y =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   171
    (case x of
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   172
      PExpr1 (PCnst c) \<Rightarrow>
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   173
        if c = 0 then y
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   174
        else
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   175
          (case y of
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   176
            PExpr1 (PCnst d) \<Rightarrow> PExpr1 (PCnst (c + d))
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   177
          | _ \<Rightarrow> PExpr1 (PAdd x y))
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   178
    | _ \<Rightarrow>
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   179
      (case y of
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   180
        PExpr1 (PCnst d) \<Rightarrow>
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   181
          if d = 0 then x
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   182
          else PExpr1 (PAdd x y)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   183
      | _ \<Rightarrow> PExpr1 (PAdd x y)))"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   184
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   185
lemma (in field) npeadd_correct:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   186
  "in_carrier xs \<Longrightarrow> peval xs (npeadd e1 e2) = peval xs (PExpr1 (PAdd e1 e2))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   187
  by (cases e1 e2 rule: pexpr_cases2) (simp_all add: npeadd_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   188
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   189
definition npesub :: "pexpr \<Rightarrow> pexpr \<Rightarrow> pexpr"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   190
  where "npesub x y =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   191
    (case y of
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   192
      PExpr1 (PCnst d) \<Rightarrow>
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   193
        if d = 0 then x
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   194
        else
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   195
          (case x of
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   196
            PExpr1 (PCnst c) \<Rightarrow> PExpr1 (PCnst (c - d))
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   197
          | _ \<Rightarrow> PExpr1 (PSub x y))
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   198
    | _ \<Rightarrow>
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   199
      (case x of
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   200
        PExpr1 (PCnst c) \<Rightarrow>
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   201
          if c = 0 then PExpr1 (PNeg y)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   202
          else PExpr1 (PSub x y)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   203
      | _ \<Rightarrow> PExpr1 (PSub x y)))"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   204
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   205
lemma (in field) npesub_correct:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   206
  "in_carrier xs \<Longrightarrow> peval xs (npesub e1 e2) = peval xs (PExpr1 (PSub e1 e2))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   207
  by (cases e1 e2 rule: pexpr_cases2) (simp_all add: npesub_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   208
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   209
definition npeneg :: "pexpr \<Rightarrow> pexpr"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   210
  where "npeneg e =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   211
    (case e of
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   212
      PExpr1 (PCnst c) \<Rightarrow> PExpr1 (PCnst (- c))
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   213
    | _ \<Rightarrow> PExpr1 (PNeg e))"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   214
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   215
lemma (in field) npeneg_correct: "peval xs (npeneg e) = peval xs (PExpr1 (PNeg e))"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   216
  by (cases e rule: pexpr_cases) (simp_all add: npeneg_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   217
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   218
lemma option_pair_cases [case_names None Some]:
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   219
  obtains (None) "x = None" | (Some) p q where "x = Some (p, q)"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   220
proof (cases x)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   221
  case None
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   222
  then show ?thesis by (rule that)
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   223
next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   224
  case (Some r)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   225
  then show ?thesis
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   226
    by (cases r, simp) (rule that)
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   227
qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   228
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   229
fun isin :: "pexpr \<Rightarrow> nat \<Rightarrow> pexpr \<Rightarrow> nat \<Rightarrow> (nat \<times> pexpr) option"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   230
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   231
    "isin e n (PExpr2 (PMul e1 e2)) m =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   232
      (case isin e n e1 m of
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   233
        Some (k, e3) \<Rightarrow>
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   234
          if k = 0 then Some (0, npemul e3 (npepow e2 m))
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   235
          else
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   236
            (case isin e k e2 m of
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   237
              Some (l, e4) \<Rightarrow> Some (l, npemul e3 e4)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   238
            | None \<Rightarrow> Some (k, npemul e3 (npepow e2 m)))
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   239
      | None \<Rightarrow>
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   240
          (case isin e n e2 m of
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   241
            Some (k, e3) \<Rightarrow> Some (k, npemul (npepow e1 m) e3)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   242
          | None \<Rightarrow> None))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   243
  | "isin e n (PExpr2 (PPow e' k)) m =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   244
      (if k = 0 then None else isin e n e' (k * m))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   245
  | "isin (PExpr1 e) n (PExpr1 e') m =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   246
      (if e = e' then
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   247
        if n \<ge> m then Some (n - m, PExpr1 (PCnst 1))
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   248
        else Some (0, npepow (PExpr1 e) (m - n))
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   249
       else None)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   250
  | "isin (PExpr2 e) n (PExpr1 e') m = None"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   251
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   252
lemma (in field) isin_correct:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   253
  assumes "in_carrier xs"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   254
    and "isin e n e' m = Some (p, e'')"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   255
  shows "peval xs (PExpr2 (PPow e' m)) = peval xs (PExpr2 (PMul (PExpr2 (PPow e (n - p))) e''))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   256
    and "p \<le> n"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   257
  using assms
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   258
  by (induct e n e' m arbitrary: p e'' rule: isin.induct)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   259
    (force
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   260
       simp add:
70019
095dce9892e8 A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   261
         nat_pow_distrib nat_pow_pow nat_pow_mult m_ac
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   262
         npemul_correct npepow_correct
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   263
       split: option.split_asm prod.split_asm if_split_asm)+
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   264
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   265
lemma (in field) isin_correct':
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   266
  "in_carrier xs \<Longrightarrow> isin e n e' 1 = Some (p, e'') \<Longrightarrow>
67341
df79ef3b3a41 Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents: 67123
diff changeset
   267
    peval xs e' = peval xs e [^] (n - p) \<otimes> peval xs e''"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   268
  "in_carrier xs \<Longrightarrow> isin e n e' 1 = Some (p, e'') \<Longrightarrow> p \<le> n"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   269
  using isin_correct [where m = 1] by simp_all
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   270
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   271
fun split_aux :: "pexpr \<Rightarrow> nat \<Rightarrow> pexpr \<Rightarrow> pexpr \<times> pexpr \<times> pexpr"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   272
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   273
    "split_aux (PExpr2 (PMul e1 e2)) n e =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   274
       (let
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   275
          (left1, common1, right1) = split_aux e1 n e;
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   276
          (left2, common2, right2) = split_aux e2 n right1
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   277
        in (npemul left1 left2, npemul common1 common2, right2))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   278
  | "split_aux (PExpr2 (PPow e' m)) n e =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   279
       (if m = 0 then (PExpr1 (PCnst 1), PExpr1 (PCnst 1), e)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   280
        else split_aux e' (m * n) e)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   281
  | "split_aux (PExpr1 e') n e =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   282
       (case isin (PExpr1 e') n e 1 of
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   283
          Some (m, e'') \<Rightarrow>
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   284
            (if m = 0 then (PExpr1 (PCnst 1), npepow (PExpr1 e') n, e'')
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   285
             else (npepow (PExpr1 e') m, npepow (PExpr1 e') (n - m), e''))
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   286
        | None \<Rightarrow> (npepow (PExpr1 e') n, PExpr1 (PCnst 1), e))"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   287
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   288
hide_const Left Right  (* FIXME !? *)
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   289
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   290
abbreviation Left :: "pexpr \<Rightarrow> pexpr \<Rightarrow> pexpr"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   291
  where "Left e1 e2 \<equiv> fst (split_aux e1 (Suc 0) e2)"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   292
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   293
abbreviation Common :: "pexpr \<Rightarrow> pexpr \<Rightarrow> pexpr"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   294
  where "Common e1 e2 \<equiv> fst (snd (split_aux e1 (Suc 0) e2))"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   295
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   296
abbreviation Right :: "pexpr \<Rightarrow> pexpr \<Rightarrow> pexpr"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   297
  where "Right e1 e2 \<equiv> snd (snd (split_aux e1 (Suc 0) e2))"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   298
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   299
lemma split_aux_induct [case_names 1 2 3]:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   300
  assumes I1: "\<And>e1 e2 n e. P e1 n e \<Longrightarrow> P e2 n (snd (snd (split_aux e1 n e))) \<Longrightarrow>
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   301
    P (PExpr2 (PMul e1 e2)) n e"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   302
  and I2: "\<And>e' m n e. (m \<noteq> 0 \<Longrightarrow> P e' (m * n) e) \<Longrightarrow> P (PExpr2 (PPow e' m)) n e"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   303
  and I3: "\<And>e' n e. P (PExpr1 e') n e"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   304
  shows "P x y z"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   305
proof (induct x y z rule: split_aux.induct)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   306
  case 1
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   307
  from 1(1) 1(2) [OF refl prod.collapse prod.collapse]
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   308
  show ?case by (rule I1)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   309
next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   310
  case 2
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   311
  then show ?case by (rule I2)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   312
next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   313
  case 3
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   314
  then show ?case by (rule I3)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   315
qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   316
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   317
lemma (in field) split_aux_correct:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   318
  "in_carrier xs \<Longrightarrow>
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   319
    peval xs (PExpr2 (PPow e\<^sub>1 n)) =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   320
    peval xs (PExpr2 (PMul (fst (split_aux e\<^sub>1 n e\<^sub>2)) (fst (snd (split_aux e\<^sub>1 n e\<^sub>2)))))"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   321
  "in_carrier xs \<Longrightarrow>
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   322
    peval xs e\<^sub>2 =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   323
    peval xs (PExpr2 (PMul (snd (snd (split_aux e\<^sub>1 n e\<^sub>2))) (fst (snd (split_aux e\<^sub>1 n e\<^sub>2)))))"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   324
  by (induct e\<^sub>1 n e\<^sub>2 rule: split_aux_induct)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   325
    (auto simp add: split_beta
70019
095dce9892e8 A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   326
       nat_pow_distrib nat_pow_pow nat_pow_mult m_ac
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   327
       npemul_correct npepow_correct isin_correct'
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   328
       split: option.split)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   329
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   330
lemma (in field) split_aux_correct':
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   331
  "in_carrier xs \<Longrightarrow>
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   332
    peval xs e\<^sub>1 = peval xs (Left e\<^sub>1 e\<^sub>2) \<otimes> peval xs (Common e\<^sub>1 e\<^sub>2)"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   333
  "in_carrier xs \<Longrightarrow>
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   334
    peval xs e\<^sub>2 = peval xs (Right e\<^sub>1 e\<^sub>2) \<otimes> peval xs (Common e\<^sub>1 e\<^sub>2)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   335
  using split_aux_correct [where n = 1] by simp_all
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   336
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   337
fun fnorm :: "fexpr \<Rightarrow> pexpr \<times> pexpr \<times> pexpr list"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   338
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   339
    "fnorm (FCnst c) = (PExpr1 (PCnst c), PExpr1 (PCnst 1), [])"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   340
  | "fnorm (FVar n) = (PExpr1 (PVar n), PExpr1 (PCnst 1), [])"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   341
  | "fnorm (FAdd e1 e2) =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   342
       (let
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   343
          (xn, xd, xc) = fnorm e1;
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   344
          (yn, yd, yc) = fnorm e2;
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   345
          (left, common, right) = split_aux xd 1 yd
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   346
        in
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   347
          (npeadd (npemul xn right) (npemul yn left),
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   348
           npemul left (npemul right common),
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   349
           List.union xc yc))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   350
  | "fnorm (FSub e1 e2) =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   351
       (let
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   352
          (xn, xd, xc) = fnorm e1;
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   353
          (yn, yd, yc) = fnorm e2;
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   354
          (left, common, right) = split_aux xd 1 yd
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   355
        in
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   356
          (npesub (npemul xn right) (npemul yn left),
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   357
           npemul left (npemul right common),
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   358
           List.union xc yc))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   359
  | "fnorm (FMul e1 e2) =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   360
       (let
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   361
          (xn, xd, xc) = fnorm e1;
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   362
          (yn, yd, yc) = fnorm e2;
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   363
          (left1, common1, right1) = split_aux xn 1 yd;
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   364
          (left2, common2, right2) = split_aux yn 1 xd
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   365
        in
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   366
          (npemul left1 left2,
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   367
           npemul right2 right1,
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   368
           List.union xc yc))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   369
  | "fnorm (FNeg e) =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   370
       (let (n, d, c) = fnorm e
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   371
        in (npeneg n, d, c))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   372
  | "fnorm (FDiv e1 e2) =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   373
       (let
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   374
          (xn, xd, xc) = fnorm e1;
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   375
          (yn, yd, yc) = fnorm e2;
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   376
          (left1, common1, right1) = split_aux xn 1 yn;
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   377
          (left2, common2, right2) = split_aux xd 1 yd
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   378
        in
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   379
          (npemul left1 right2,
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   380
           npemul left2 right1,
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   381
           List.insert yn (List.union xc yc)))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   382
  | "fnorm (FPow e m) =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   383
       (let (n, d, c) = fnorm e
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   384
        in (npepow n m, npepow d m, c))"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   385
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   386
abbreviation Num :: "fexpr \<Rightarrow> pexpr"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   387
  where "Num e \<equiv> fst (fnorm e)"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   388
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   389
abbreviation Denom :: "fexpr \<Rightarrow> pexpr"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   390
  where "Denom e \<equiv> fst (snd (fnorm e))"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   391
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   392
abbreviation Cond :: "fexpr \<Rightarrow> pexpr list"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   393
  where "Cond e \<equiv> snd (snd (fnorm e))"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   394
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   395
primrec (in field) nonzero :: "'a list \<Rightarrow> pexpr list \<Rightarrow> bool"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   396
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   397
    "nonzero xs [] \<longleftrightarrow> True"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   398
  | "nonzero xs (p # ps) \<longleftrightarrow> peval xs p \<noteq> \<zero> \<and> nonzero xs ps"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   399
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   400
lemma (in field) nonzero_singleton: "nonzero xs [p] = (peval xs p \<noteq> \<zero>)"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   401
  by simp
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   402
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   403
lemma (in field) nonzero_append: "nonzero xs (ps @ qs) = (nonzero xs ps \<and> nonzero xs qs)"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   404
  by (induct ps) simp_all
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   405
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   406
lemma (in field) nonzero_idempotent:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   407
  "p \<in> set ps \<Longrightarrow> (peval xs p \<noteq> \<zero> \<and> nonzero xs ps) = nonzero xs ps"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   408
  by (induct ps) auto
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   409
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   410
lemma (in field) nonzero_insert:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   411
  "nonzero xs (List.insert p ps) = (peval xs p \<noteq> \<zero> \<and> nonzero xs ps)"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   412
  by (simp add: List.insert_def nonzero_idempotent)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   413
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   414
lemma (in field) nonzero_union:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   415
  "nonzero xs (List.union ps qs) = (nonzero xs ps \<and> nonzero xs qs)"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   416
  by (induct ps rule: rev_induct)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   417
    (auto simp add: List.union_def nonzero_insert nonzero_append)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   418
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   419
lemma (in field) fnorm_correct:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   420
  assumes "in_carrier xs"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   421
    and "nonzero xs (Cond e)"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   422
  shows "feval xs e = peval xs (Num e) \<oslash> peval xs (Denom e)"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   423
    and "peval xs (Denom e) \<noteq> \<zero>"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   424
  using assms
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   425
proof (induct e)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   426
  case (FCnst c)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   427
  {
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   428
    case 1
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   429
    show ?case by simp
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   430
  next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   431
    case 2
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   432
    show ?case by simp
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   433
  }
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   434
next
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   435
  case (FVar n)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   436
  {
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   437
    case 1
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   438
    then show ?case by simp
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   439
  next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   440
    case 2
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   441
    show ?case by simp
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   442
  }
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   443
next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   444
  case (FAdd e1 e2)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   445
  note split = split_aux_correct' [where xs=xs and e\<^sub>1 = "Denom e1" and e\<^sub>2 = "Denom e2"]
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   446
  {
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   447
    case 1
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   448
    let ?left = "peval xs (Left (Denom e1) (Denom e2))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   449
    let ?common = "peval xs (Common (Denom e1) (Denom e2))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   450
    let ?right = "peval xs (Right (Denom e1) (Denom e2))"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   451
    from 1 FAdd have "feval xs (FAdd e1 e2) =
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   452
      (?common \<otimes> (peval xs (Num e1) \<otimes> ?right \<oplus> peval xs (Num e2) \<otimes> ?left)) \<oslash>
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   453
      (?common \<otimes> (?left \<otimes> (?right \<otimes> ?common)))"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   454
      by (simp add: split_beta split nonzero_union add_frac_eq r_distr m_ac)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   455
    also from 1 FAdd have "\<dots> = peval xs (Num (FAdd e1 e2)) \<oslash> peval xs (Denom (FAdd e1 e2))"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   456
      by (simp add: split_beta split nonzero_union npeadd_correct npemul_correct integral_iff)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   457
    finally show ?case .
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   458
  next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   459
    case 2
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   460
    with FAdd show ?case
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   461
      by (simp add: split_beta split nonzero_union npemul_correct integral_iff)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   462
  }
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   463
next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   464
  case (FSub e1 e2)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   465
  note split = split_aux_correct' [where xs=xs and e\<^sub>1 = "Denom e1" and e\<^sub>2 = "Denom e2"]
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   466
  {
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   467
    case 1
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   468
    let ?left = "peval xs (Left (Denom e1) (Denom e2))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   469
    let ?common = "peval xs (Common (Denom e1) (Denom e2))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   470
    let ?right = "peval xs (Right (Denom e1) (Denom e2))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   471
    from 1 FSub
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   472
    have "feval xs (FSub e1 e2) =
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   473
      (?common \<otimes> (peval xs (Num e1) \<otimes> ?right \<ominus> peval xs (Num e2) \<otimes> ?left)) \<oslash>
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   474
      (?common \<otimes> (?left \<otimes> (?right \<otimes> ?common)))"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   475
      by (simp add: split_beta split nonzero_union diff_frac_eq r_diff_distr m_ac)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   476
    also from 1 FSub have "\<dots> = peval xs (Num (FSub e1 e2)) \<oslash> peval xs (Denom (FSub e1 e2))"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   477
      by (simp add: split_beta split nonzero_union npesub_correct npemul_correct integral_iff)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   478
    finally show ?case .
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   479
  next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   480
    case 2
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   481
    with FSub show ?case
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   482
      by (simp add: split_beta split nonzero_union npemul_correct integral_iff)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   483
  }
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   484
next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   485
  case (FMul e1 e2)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   486
  note split =
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   487
    split_aux_correct' [where xs=xs and e\<^sub>1 = "Num e1" and e\<^sub>2 = "Denom e2"]
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   488
    split_aux_correct' [where xs=xs and e\<^sub>1 = "Num e2" and e\<^sub>2 = "Denom e1"]
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   489
  {
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   490
    case 1
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   491
    let ?left\<^sub>1 = "peval xs (Left (Num e1) (Denom e2))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   492
    let ?common\<^sub>1 = "peval xs (Common (Num e1) (Denom e2))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   493
    let ?right\<^sub>1 = "peval xs (Right (Num e1) (Denom e2))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   494
    let ?left\<^sub>2 = "peval xs (Left (Num e2) (Denom e1))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   495
    let ?common\<^sub>2 = "peval xs (Common (Num e2) (Denom e1))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   496
    let ?right\<^sub>2 = "peval xs (Right (Num e2) (Denom e1))"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   497
    from 1 FMul have "feval xs (FMul e1 e2) =
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   498
      ((?common\<^sub>1 \<otimes> ?common\<^sub>2) \<otimes> (?left\<^sub>1 \<otimes> ?left\<^sub>2)) \<oslash>
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   499
      ((?common\<^sub>1 \<otimes> ?common\<^sub>2) \<otimes> (?right\<^sub>2 \<otimes> ?right\<^sub>1))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   500
      by (simp add: split_beta split nonzero_union
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   501
        nonzero_divide_divide_eq_left m_ac)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   502
    also from 1 FMul have "\<dots> = peval xs (Num (FMul e1 e2)) \<oslash> peval xs (Denom (FMul e1 e2))"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   503
      by (simp add: split_beta split nonzero_union npemul_correct integral_iff)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   504
    finally show ?case .
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   505
  next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   506
    case 2
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   507
    with FMul show ?case
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   508
      by (simp add: split_beta split nonzero_union npemul_correct integral_iff)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   509
  }
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   510
next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   511
  case (FNeg e)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   512
  {
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   513
    case 1
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   514
    with FNeg show ?case
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   515
      by (simp add: split_beta npeneg_correct)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   516
  next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   517
    case 2
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   518
    with FNeg show ?case
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   519
      by (simp add: split_beta)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   520
  }
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   521
next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   522
  case (FDiv e1 e2)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   523
  note split =
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   524
    split_aux_correct' [where xs=xs and e\<^sub>1 = "Num e1" and e\<^sub>2 = "Num e2"]
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   525
    split_aux_correct' [where xs=xs and e\<^sub>1 = "Denom e1" and e\<^sub>2 = "Denom e2"]
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   526
  {
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   527
    case 1
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   528
    let ?left\<^sub>1 = "peval xs (Left (Num e1) (Num e2))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   529
    let ?common\<^sub>1 = "peval xs (Common (Num e1) (Num e2))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   530
    let ?right\<^sub>1 = "peval xs (Right (Num e1) (Num e2))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   531
    let ?left\<^sub>2 = "peval xs (Left (Denom e1) (Denom e2))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   532
    let ?common\<^sub>2 = "peval xs (Common (Denom e1) (Denom e2))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   533
    let ?right\<^sub>2 = "peval xs (Right (Denom e1) (Denom e2))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   534
    from 1 FDiv
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   535
    have "feval xs (FDiv e1 e2) =
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   536
      ((?common\<^sub>1 \<otimes> ?common\<^sub>2) \<otimes> (?left\<^sub>1 \<otimes> ?right\<^sub>2)) \<oslash>
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   537
      ((?common\<^sub>1 \<otimes> ?common\<^sub>2) \<otimes> (?left\<^sub>2 \<otimes> ?right\<^sub>1))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   538
      by (simp add: split_beta split nonzero_union nonzero_insert
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   539
        nonzero_divide_divide_eq m_ac)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   540
    also from 1 FDiv have "\<dots> = peval xs (Num (FDiv e1 e2)) \<oslash> peval xs (Denom (FDiv e1 e2))"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   541
      by (simp add: split_beta split nonzero_union nonzero_insert npemul_correct integral_iff)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   542
    finally show ?case .
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   543
  next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   544
    case 2
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   545
    with FDiv show ?case
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   546
      by (simp add: split_beta split nonzero_union nonzero_insert npemul_correct integral_iff)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   547
  }
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   548
next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   549
  case (FPow e n)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   550
  {
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   551
    case 1
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   552
    with FPow show ?case
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   553
      by (simp add: split_beta nonzero_power_divide npepow_correct)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   554
  next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   555
    case 2
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   556
    with FPow show ?case
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   557
      by (simp add: split_beta npepow_correct)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   558
  }
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   559
qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   560
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   561
lemma (in field) feval_eq0:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   562
  assumes "in_carrier xs"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   563
    and "fnorm e = (n, d, c)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   564
    and "nonzero xs c"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   565
    and "peval xs n = \<zero>"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   566
  shows "feval xs e = \<zero>"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   567
  using assms fnorm_correct [of xs e] by simp
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   568
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   569
lemma (in field) fexpr_in_carrier:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   570
  assumes "in_carrier xs"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   571
    and "nonzero xs (Cond e)"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   572
  shows "feval xs e \<in> carrier R"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   573
  using assms
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   574
proof (induct e)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   575
  case (FDiv e1 e2)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   576
  then have "feval xs e1 \<in> carrier R" "feval xs e2 \<in> carrier R"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   577
    "peval xs (Num e2) \<noteq> \<zero>" "nonzero xs (Cond e2)"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   578
    by (simp_all add: nonzero_union nonzero_insert split: prod.split_asm)
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   579
  from \<open>in_carrier xs\<close> \<open>nonzero xs (Cond e2)\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   580
  have "feval xs e2 = peval xs (Num e2) \<oslash> peval xs (Denom e2)"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   581
    by (rule fnorm_correct)
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   582
  moreover from \<open>in_carrier xs\<close> \<open>nonzero xs (Cond e2)\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   583
  have "peval xs (Denom e2) \<noteq> \<zero>" by (rule fnorm_correct)
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   584
  ultimately have "feval xs e2 \<noteq> \<zero>" using \<open>peval xs (Num e2) \<noteq> \<zero>\<close> \<open>in_carrier xs\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   585
    by (simp add: divide_eq_0_iff)
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   586
  with \<open>feval xs e1 \<in> carrier R\<close> \<open>feval xs e2 \<in> carrier R\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   587
  show ?case by simp
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   588
qed (simp_all add: nonzero_union split: prod.split_asm)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   589
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   590
lemma (in field) feval_eq:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   591
  assumes "in_carrier xs"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   592
    and "fnorm (FSub e e') = (n, d, c)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   593
    and "nonzero xs c"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   594
  shows "(feval xs e = feval xs e') = (peval xs n = \<zero>)"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   595
proof -
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   596
  from assms have "nonzero xs (Cond e)" "nonzero xs (Cond e')"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   597
    by (auto simp add: nonzero_union split: prod.split_asm)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   598
  with assms fnorm_correct [of xs "FSub e e'"]
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   599
  have "feval xs e \<ominus> feval xs e' = peval xs n \<oslash> peval xs d"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   600
    "peval xs d \<noteq> \<zero>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   601
    by simp_all
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   602
  show ?thesis
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   603
  proof
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   604
    assume "feval xs e = feval xs e'"
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   605
    with \<open>feval xs e \<ominus> feval xs e' = peval xs n \<oslash> peval xs d\<close>
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   606
      \<open>in_carrier xs\<close> \<open>nonzero xs (Cond e')\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   607
    have "peval xs n \<oslash> peval xs d = \<zero>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   608
      by (simp add: fexpr_in_carrier minus_eq r_neg)
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   609
    with \<open>peval xs d \<noteq> \<zero>\<close> \<open>in_carrier xs\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   610
    show "peval xs n = \<zero>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   611
      by (simp add: divide_eq_0_iff)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   612
  next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   613
    assume "peval xs n = \<zero>"
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   614
    with \<open>feval xs e \<ominus> feval xs e' = peval xs n \<oslash> peval xs d\<close> \<open>peval xs d \<noteq> \<zero>\<close>
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   615
      \<open>nonzero xs (Cond e)\<close> \<open>nonzero xs (Cond e')\<close> \<open>in_carrier xs\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   616
    show "feval xs e = feval xs e'"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   617
      by (simp add: eq_diff0 fexpr_in_carrier)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   618
  qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   619
qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   620
64999
f8f76a501d25 elaborated examples for computations
haftmann
parents: 64998
diff changeset
   621
ML \<open>
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   622
val term_of_nat = HOLogic.mk_number \<^Type>\<open>nat\<close> o @{code integer_of_nat};
64999
f8f76a501d25 elaborated examples for computations
haftmann
parents: 64998
diff changeset
   623
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   624
val term_of_int = HOLogic.mk_number \<^Type>\<open>int\<close> o @{code integer_of_int};
64999
f8f76a501d25 elaborated examples for computations
haftmann
parents: 64998
diff changeset
   625
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   626
fun term_of_pexpr (@{code PExpr1} x) = \<^Const>\<open>PExpr1\<close> $ term_of_pexpr1 x
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   627
  | term_of_pexpr (@{code PExpr2} x) = \<^Const>\<open>PExpr2\<close> $ term_of_pexpr2 x
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   628
and term_of_pexpr1 (@{code PCnst} k) = \<^Const>\<open>PCnst\<close> $ term_of_int k
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   629
  | term_of_pexpr1 (@{code PVar} n) = \<^Const>\<open>PVar\<close> $ term_of_nat n
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   630
  | term_of_pexpr1 (@{code PAdd} (x, y)) = \<^Const>\<open>PAdd\<close> $ term_of_pexpr x $ term_of_pexpr y
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   631
  | term_of_pexpr1 (@{code PSub} (x, y)) = \<^Const>\<open>PSub\<close> $ term_of_pexpr x $ term_of_pexpr y
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   632
  | term_of_pexpr1 (@{code PNeg} x) = \<^Const>\<open>PNeg\<close> $ term_of_pexpr x
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   633
and term_of_pexpr2 (@{code PMul} (x, y)) = \<^Const>\<open>PMul\<close> $ term_of_pexpr x $ term_of_pexpr y
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   634
  | term_of_pexpr2 (@{code PPow} (x, n)) = \<^Const>\<open>PPow\<close> $ term_of_pexpr x $ term_of_nat n
64999
f8f76a501d25 elaborated examples for computations
haftmann
parents: 64998
diff changeset
   635
f8f76a501d25 elaborated examples for computations
haftmann
parents: 64998
diff changeset
   636
fun term_of_result (x, (y, zs)) =
f8f76a501d25 elaborated examples for computations
haftmann
parents: 64998
diff changeset
   637
  HOLogic.mk_prod (term_of_pexpr x, HOLogic.mk_prod
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   638
    (term_of_pexpr y, HOLogic.mk_list \<^Type>\<open>pexpr\<close> (map term_of_pexpr zs)));
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   639
64999
f8f76a501d25 elaborated examples for computations
haftmann
parents: 64998
diff changeset
   640
local
f8f76a501d25 elaborated examples for computations
haftmann
parents: 64998
diff changeset
   641
74569
f4613ca298e6 more antiquotations;
wenzelm
parents: 74561
diff changeset
   642
fun fnorm (ctxt, ct, t) =
74570
7625b5d7cfe2 clarified name, syntax, messages;
wenzelm
parents: 74569
diff changeset
   643
  \<^instantiate>\<open>x = ct and y = \<open>Thm.cterm_of ctxt t\<close>
74569
f4613ca298e6 more antiquotations;
wenzelm
parents: 74561
diff changeset
   644
    in cterm \<open>x \<equiv> y\<close> for x y :: \<open>pexpr \<times> pexpr \<times> pexpr list\<close>\<close>;
64999
f8f76a501d25 elaborated examples for computations
haftmann
parents: 64998
diff changeset
   645
f8f76a501d25 elaborated examples for computations
haftmann
parents: 64998
diff changeset
   646
val (_, raw_fnorm_oracle) = Context.>>> (Context.map_theory_result
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   647
  (Thm.add_oracle (\<^binding>\<open>fnorm\<close>, fnorm)));
64999
f8f76a501d25 elaborated examples for computations
haftmann
parents: 64998
diff changeset
   648
f8f76a501d25 elaborated examples for computations
haftmann
parents: 64998
diff changeset
   649
fun fnorm_oracle ctxt ct t = raw_fnorm_oracle (ctxt, ct, t);
f8f76a501d25 elaborated examples for computations
haftmann
parents: 64998
diff changeset
   650
f8f76a501d25 elaborated examples for computations
haftmann
parents: 64998
diff changeset
   651
in
f8f76a501d25 elaborated examples for computations
haftmann
parents: 64998
diff changeset
   652
f8f76a501d25 elaborated examples for computations
haftmann
parents: 64998
diff changeset
   653
val cv = @{computation_conv "pexpr \<times> pexpr \<times> pexpr list"
f8f76a501d25 elaborated examples for computations
haftmann
parents: 64998
diff changeset
   654
  terms: fnorm nat_of_integer Code_Target_Nat.natural
f8f76a501d25 elaborated examples for computations
haftmann
parents: 64998
diff changeset
   655
    "0::nat" "1::nat" "2::nat" "3::nat"
f8f76a501d25 elaborated examples for computations
haftmann
parents: 64998
diff changeset
   656
    "0::int" "1::int" "2::int" "3::int" "-1::int"
f8f76a501d25 elaborated examples for computations
haftmann
parents: 64998
diff changeset
   657
  datatypes: fexpr int integer num}
65043
fd753468786f explicit dynamic context for gap-bridging function
haftmann
parents: 64999
diff changeset
   658
  (fn ctxt => fn result => fn ct => fnorm_oracle ctxt ct (term_of_result result))
64999
f8f76a501d25 elaborated examples for computations
haftmann
parents: 64998
diff changeset
   659
f8f76a501d25 elaborated examples for computations
haftmann
parents: 64998
diff changeset
   660
end
f8f76a501d25 elaborated examples for computations
haftmann
parents: 64998
diff changeset
   661
\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   662
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   663
ML \<open>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   664
signature FIELD_TAC =
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   665
sig
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   666
  structure Field_Simps:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   667
  sig
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   668
    type T
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   669
    val get: Context.generic -> T
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   670
    val put: T -> Context.generic -> Context.generic
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   671
    val map: (T -> T) -> Context.generic -> Context.generic
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   672
  end
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   673
  val eq_field_simps:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   674
    (term * (thm list * thm list * thm list * thm * thm)) *
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   675
    (term * (thm list * thm list * thm list * thm * thm)) -> bool
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   676
  val field_tac: bool -> Proof.context -> int -> tactic
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   677
end
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   678
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   679
structure Field_Tac : FIELD_TAC =
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   680
struct
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   681
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   682
open Ring_Tac;
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   683
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   684
fun field_struct \<^Const_>\<open>Ring.ring.add _ _ for R _ _\<close> = SOME R
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   685
  | field_struct \<^Const_>\<open>Ring.a_minus _ _ for R _ _\<close> = SOME R
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   686
  | field_struct \<^Const_>\<open>Group.monoid.mult _ _ for R _ _\<close> = SOME R
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   687
  | field_struct \<^Const_>\<open>Ring.a_inv _ _ for R _\<close> = SOME R
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   688
  | field_struct \<^Const_>\<open>Group.pow _ _ _ for R _ _\<close> = SOME R
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   689
  | field_struct \<^Const_>\<open>Algebra_Aux.m_div _ _for R _ _\<close> = SOME R
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   690
  | field_struct \<^Const_>\<open>Ring.ring.zero _ _ for R\<close> = SOME R
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   691
  | field_struct \<^Const_>\<open>Group.monoid.one _ _ for R\<close> = SOME R
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   692
  | field_struct \<^Const_>\<open>Algebra_Aux.of_integer _ _ for R _\<close> = SOME R
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   693
  | field_struct _ = NONE;
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   694
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   695
fun reif_fexpr vs \<^Const_>\<open>Ring.ring.add _ _ for _ a b\<close> =
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   696
      \<^Const>\<open>FAdd for \<open>reif_fexpr vs a\<close> \<open>reif_fexpr vs b\<close>\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   697
  | reif_fexpr vs \<^Const_>\<open>Ring.a_minus _ _ for _ a b\<close> =
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   698
      \<^Const>\<open>FSub for \<open>reif_fexpr vs a\<close> \<open>reif_fexpr vs b\<close>\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   699
  | reif_fexpr vs \<^Const_>\<open>Group.monoid.mult _ _ for _ a b\<close> =
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   700
      \<^Const>\<open>FMul for \<open>reif_fexpr vs a\<close> \<open>reif_fexpr vs b\<close>\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   701
  | reif_fexpr vs \<^Const_>\<open>Ring.a_inv _ _ for _ a\<close> =
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   702
      \<^Const>\<open>FNeg for \<open>reif_fexpr vs a\<close>\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   703
  | reif_fexpr vs \<^Const>\<open>Group.pow _ _ _ for _ a n\<close> =
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   704
      \<^Const>\<open>FPow for \<open>reif_fexpr vs a\<close> n\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   705
  | reif_fexpr vs \<^Const_>\<open>Algebra_Aux.m_div _ _ for _ a b\<close> =
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   706
      \<^Const>\<open>FDiv for \<open>reif_fexpr vs a\<close> \<open>reif_fexpr vs b\<close>\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   707
  | reif_fexpr vs (Free x) =
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   708
      \<^Const>\<open>FVar for \<open>HOLogic.mk_number HOLogic.natT (find_index (equal x) vs)\<close>\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   709
  | reif_fexpr vs \<^Const_>\<open>Ring.ring.zero _ _ for _\<close> = \<^term>\<open>FCnst 0\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   710
  | reif_fexpr vs \<^Const_>\<open>Group.monoid.one _ _ for _\<close> = \<^term>\<open>FCnst 1\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   711
  | reif_fexpr vs \<^Const_>\<open>Algebra_Aux.of_integer _ _ for _ n\<close> = \<^Const>\<open>FCnst for n\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   712
  | reif_fexpr _ _ = error "reif_fexpr: bad expression";
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   713
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   714
fun reif_fexpr' vs \<^Const_>\<open>plus _ for a b\<close> = \<^Const>\<open>FAdd for \<open>reif_fexpr' vs a\<close> \<open>reif_fexpr' vs b\<close>\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   715
  | reif_fexpr' vs \<^Const_>\<open>minus _ for a b\<close> = \<^Const>\<open>FSub for \<open>reif_fexpr' vs a\<close> \<open>reif_fexpr' vs b\<close>\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   716
  | reif_fexpr' vs \<^Const_>\<open>times _ for a b\<close> = \<^Const>\<open>FMul for \<open>reif_fexpr' vs a\<close> \<open>reif_fexpr' vs b\<close>\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   717
  | reif_fexpr' vs \<^Const_>\<open>uminus _ for a\<close> = \<^Const>\<open>FNeg for \<open>reif_fexpr' vs a\<close>\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   718
  | reif_fexpr' vs \<^Const_>\<open>power _ for a n\<close> = \<^Const>\<open>FPow for \<open>reif_fexpr' vs a\<close> n\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   719
  | reif_fexpr' vs \<^Const_>\<open>divide _ for a b\<close> = \<^Const>\<open>FDiv for \<open>reif_fexpr' vs a\<close> \<open>reif_fexpr' vs b\<close>\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   720
  | reif_fexpr' vs (Free x) =
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   721
      \<^Const>\<open>FVar for \<open>HOLogic.mk_number HOLogic.natT (find_index (equal x) vs)\<close>\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   722
  | reif_fexpr' vs \<^Const_>\<open>zero_class.zero _\<close> = \<^term>\<open>FCnst 0\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   723
  | reif_fexpr' vs \<^Const_>\<open>one_class.one _\<close> = \<^term>\<open>FCnst 1\<close>
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   724
  | reif_fexpr' vs \<^Const_>\<open>numeral _ for b\<close> = \<^Const>\<open>FCnst for \<^Const>\<open>numeral \<^Type>\<open>int\<close> for b\<close>\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   725
  | reif_fexpr' _ _ = error "reif_fexpr: bad expression";
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   726
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   727
fun eq_field_simps
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   728
  ((t, (ths1, ths2, ths3, th4, th)),
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   729
   (t', (ths1', ths2', ths3', th4', th'))) =
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   730
    t aconv t' andalso
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   731
    eq_list Thm.eq_thm (ths1, ths1') andalso
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   732
    eq_list Thm.eq_thm (ths2, ths2') andalso
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   733
    eq_list Thm.eq_thm (ths3, ths3') andalso
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   734
    Thm.eq_thm (th4, th4') andalso
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   735
    Thm.eq_thm (th, th');
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   736
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   737
structure Field_Simps = Generic_Data
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   738
(struct
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   739
  type T = (term * (thm list * thm list * thm list * thm * thm)) Net.net
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   740
  val empty = Net.empty
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   741
  val merge = Net.merge eq_field_simps
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   742
end);
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   743
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   744
fun get_field_simps ctxt optcT t =
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   745
  (case get_matching_rules ctxt (Field_Simps.get (Context.Proof ctxt)) t of
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   746
     SOME (ths1, ths2, ths3, th4, th) =>
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   747
       let val tr =
67649
1e1782c1aedf tuned signature;
wenzelm
parents: 67341
diff changeset
   748
         Thm.transfer' ctxt #>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   749
         (case optcT of NONE => I | SOME cT => inst [cT] [] #> norm)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   750
       in (map tr ths1, map tr ths2, map tr ths3, tr th4, tr th) end
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   751
   | NONE => error "get_field_simps: lookup failed");
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   752
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   753
fun nth_el_conv (_, _, _, nth_el_Cons, _) =
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   754
  let
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   755
    val a = type_of_eqn nth_el_Cons;
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   756
    val If_conv_a = If_conv a;
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   757
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   758
    fun conv ys n = (case strip_app ys of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   759
      (\<^const_name>\<open>Cons\<close>, [x, xs]) =>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   760
        transitive'
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   761
          (inst [] [x, xs, n] nth_el_Cons)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   762
          (If_conv_a (args2 nat_eq_conv)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   763
             Thm.reflexive
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   764
             (cong2' conv Thm.reflexive (args2 nat_minus_conv))))
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   765
  in conv end;
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   766
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   767
fun feval_conv (rls as
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   768
      ([feval_simps_1, feval_simps_2, feval_simps_3,
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   769
        feval_simps_4, feval_simps_5, feval_simps_6,
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   770
        feval_simps_7, feval_simps_8, feval_simps_9,
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   771
        feval_simps_10, feval_simps_11],
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   772
       _, _, _, _)) =
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   773
  let
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   774
    val nth_el_conv' = nth_el_conv rls;
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   775
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   776
    fun conv xs x = (case strip_app x of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   777
        (\<^const_name>\<open>FCnst\<close>, [c]) => (case strip_app c of
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   778
            (\<^const_name>\<open>zero_class.zero\<close>, _) => inst [] [xs] feval_simps_9
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   779
          | (\<^const_name>\<open>one_class.one\<close>, _) => inst [] [xs] feval_simps_10
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   780
          | (\<^const_name>\<open>numeral\<close>, [n]) => inst [] [xs, n] feval_simps_11
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   781
          | _ => inst [] [xs, c] feval_simps_1)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   782
      | (\<^const_name>\<open>FVar\<close>, [n]) =>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   783
          transitive' (inst [] [xs, n] feval_simps_2) (args2 nth_el_conv')
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   784
      | (\<^const_name>\<open>FAdd\<close>, [a, b]) =>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   785
          transitive' (inst [] [xs, a, b] feval_simps_3)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   786
            (cong2 (args2 conv) (args2 conv))
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   787
      | (\<^const_name>\<open>FSub\<close>, [a, b]) =>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   788
          transitive' (inst [] [xs, a, b] feval_simps_4)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   789
            (cong2 (args2 conv) (args2 conv))
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   790
      | (\<^const_name>\<open>FMul\<close>, [a, b]) =>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   791
          transitive' (inst [] [xs, a, b] feval_simps_5)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   792
            (cong2 (args2 conv) (args2 conv))
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   793
      | (\<^const_name>\<open>FNeg\<close>, [a]) =>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   794
          transitive' (inst [] [xs, a] feval_simps_6)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   795
            (cong1 (args2 conv))
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   796
      | (\<^const_name>\<open>FDiv\<close>, [a, b]) =>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   797
          transitive' (inst [] [xs, a, b] feval_simps_7)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   798
            (cong2 (args2 conv) (args2 conv))
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   799
      | (\<^const_name>\<open>FPow\<close>, [a, n]) =>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   800
          transitive' (inst [] [xs, a, n] feval_simps_8)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   801
            (cong2 (args2 conv) Thm.reflexive))
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   802
  in conv end;
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   803
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   804
fun peval_conv (rls as
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   805
      (_,
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   806
       [peval_simps_1, peval_simps_2, peval_simps_3,
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   807
        peval_simps_4, peval_simps_5, peval_simps_6,
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   808
        peval_simps_7, peval_simps_8, peval_simps_9,
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   809
        peval_simps_10, peval_simps_11],
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   810
       _, _, _)) =
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   811
  let
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   812
    val nth_el_conv' = nth_el_conv rls;
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   813
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   814
    fun conv xs x = (case strip_app x of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   815
        (\<^const_name>\<open>PExpr1\<close>, [e]) => (case strip_app e of
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   816
            (\<^const_name>\<open>PCnst\<close>, [c]) => (case strip_numeral c of
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   817
                (\<^const_name>\<open>zero_class.zero\<close>, _) => inst [] [xs] peval_simps_8
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   818
              | (\<^const_name>\<open>one_class.one\<close>, _) => inst [] [xs] peval_simps_9
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   819
              | (\<^const_name>\<open>numeral\<close>, [n]) => inst [] [xs, n] peval_simps_10
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   820
              | (\<^const_name>\<open>uminus\<close>, [n]) => inst [] [xs, n] peval_simps_11
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   821
              | _ => inst [] [xs, c] peval_simps_1)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   822
          | (\<^const_name>\<open>PVar\<close>, [n]) =>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   823
              transitive' (inst [] [xs, n] peval_simps_2) (args2 nth_el_conv')
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   824
          | (\<^const_name>\<open>PAdd\<close>, [a, b]) =>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   825
              transitive' (inst [] [xs, a, b] peval_simps_3)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   826
                (cong2 (args2 conv) (args2 conv))
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   827
          | (\<^const_name>\<open>PSub\<close>, [a, b]) =>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   828
              transitive' (inst [] [xs, a, b] peval_simps_4)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   829
                (cong2 (args2 conv) (args2 conv))
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   830
          | (\<^const_name>\<open>PNeg\<close>, [a]) =>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   831
              transitive' (inst [] [xs, a] peval_simps_5)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   832
                (cong1 (args2 conv)))
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   833
      | (\<^const_name>\<open>PExpr2\<close>, [e]) => (case strip_app e of
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   834
            (\<^const_name>\<open>PMul\<close>, [a, b]) =>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   835
              transitive' (inst [] [xs, a, b] peval_simps_6)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   836
                (cong2 (args2 conv) (args2 conv))
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   837
          | (\<^const_name>\<open>PPow\<close>, [a, n]) =>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   838
              transitive' (inst [] [xs, a, n] peval_simps_7)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   839
                (cong2 (args2 conv) Thm.reflexive)))
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   840
  in conv end;
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   841
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   842
fun nonzero_conv (rls as
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   843
      (_, _,
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   844
       [nonzero_Nil, nonzero_Cons, nonzero_singleton],
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   845
       _, _)) =
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   846
  let
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   847
    val peval_conv' = peval_conv rls;
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   848
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   849
    fun conv xs qs = (case strip_app qs of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   850
        (\<^const_name>\<open>Nil\<close>, []) => inst [] [xs] nonzero_Nil
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   851
      | (\<^const_name>\<open>Cons\<close>, [p, ps]) => (case Thm.term_of ps of
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   852
            \<^Const_>\<open>Nil _\<close> =>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   853
              transitive' (inst [] [xs, p] nonzero_singleton)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   854
                (cong1 (cong2 (args2 peval_conv') Thm.reflexive))
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   855
          | _ => transitive' (inst [] [xs, p, ps] nonzero_Cons)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   856
              (cong2 (cong1 (cong2 (args2 peval_conv') Thm.reflexive)) (args2 conv))))
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   857
  in conv end;
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   858
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   859
fun field_tac in_prem ctxt =
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   860
  SUBGOAL (fn (g, i) =>
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   861
    let
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   862
      val (prems, concl) = Logic.strip_horn g;
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   863
      fun find_eq s = (case s of
74397
e80c4cde6064 clarified antiquotations;
wenzelm
parents: 70019
diff changeset
   864
          (_ $ \<^Const_>\<open>HOL.eq T for t u\<close>) =>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   865
            (case (field_struct t, field_struct u) of
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   866
               (SOME R, _) => SOME ((t, u), R, T, NONE, mk_in_carrier ctxt R [], reif_fexpr)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   867
             | (_, SOME R) => SOME ((t, u), R, T, NONE, mk_in_carrier ctxt R [], reif_fexpr)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   868
             | _ =>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   869
                 if Sign.of_sort (Proof_Context.theory_of ctxt) (T, \<^sort>\<open>field\<close>)
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   870
                 then SOME ((t, u), mk_ring T, T, SOME T, K @{thm in_carrier_trivial}, reif_fexpr')
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   871
                 else NONE)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   872
        | _ => NONE);
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   873
      val ((t, u), R, T, optT, mkic, reif) =
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   874
        (case get_first find_eq
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   875
           (if in_prem then prems else [concl]) of
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   876
           SOME q => q
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   877
         | NONE => error "cannot determine field");
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   878
      val rls as (_, _, _, _, feval_eq) =
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   879
        get_field_simps ctxt (Option.map (Thm.ctyp_of ctxt) optT) R;
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   880
      val xs = [] |> Term.add_frees t |> Term.add_frees u |> filter (equal T o snd);
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   881
      val cxs = Thm.cterm_of ctxt (HOLogic.mk_list T (map Free xs));
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   882
      val ce = Thm.cterm_of ctxt (reif xs t);
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   883
      val ce' = Thm.cterm_of ctxt (reif xs u);
74570
7625b5d7cfe2 clarified name, syntax, messages;
wenzelm
parents: 74569
diff changeset
   884
      val fnorm = cv ctxt \<^instantiate>\<open>e = ce and e' = ce' in cterm \<open>fnorm (FSub e e')\<close>\<close>;
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   885
      val (_, [n, dc]) = strip_app (Thm.rhs_of fnorm);
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   886
      val (_, [_, c]) = strip_app dc;
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   887
      val th =
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   888
        Conv.fconv_rule (Conv.concl_conv 1 (Conv.arg_conv
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   889
          (binop_conv
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   890
             (binop_conv
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   891
                (K (feval_conv rls cxs ce)) (K (feval_conv rls cxs ce')))
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   892
             (Conv.arg1_conv (K (peval_conv rls cxs n))))))
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   893
        ([mkic xs,
67710
cc2db3239932 added HOLogic.mk_obj_eq convenience and eliminated some clones;
wenzelm
parents: 67649
diff changeset
   894
          HOLogic.mk_obj_eq fnorm,
cc2db3239932 added HOLogic.mk_obj_eq convenience and eliminated some clones;
wenzelm
parents: 67649
diff changeset
   895
          HOLogic.mk_obj_eq (nonzero_conv rls cxs c) RS @{thm iffD2}] MRS
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   896
         feval_eq);
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   897
      val th' = Drule.rotate_prems 1
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   898
        (th RS (if in_prem then @{thm iffD1} else @{thm iffD2}));
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   899
    in
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   900
      if in_prem then
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   901
        dresolve_tac ctxt [th'] 1 THEN defer_tac 1
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   902
      else
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   903
        resolve_tac ctxt [th'] 1
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   904
    end);
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   905
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   906
end
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   907
\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   908
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   909
context field
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 65043
diff changeset
   910
begin
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   911
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   912
local_setup \<open>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   913
Local_Theory.declaration {syntax = false, pervasive = false}
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   914
  (fn phi => Field_Tac.Field_Simps.map (Ring_Tac.insert_rules Field_Tac.eq_field_simps
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
   915
    (Morphism.term phi \<^term>\<open>R\<close>,
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   916
     (Morphism.fact phi @{thms feval.simps [meta] feval_Cnst [meta]},
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   917
      Morphism.fact phi @{thms peval.simps [meta] peval_Cnst [meta]},
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   918
      Morphism.fact phi @{thms nonzero.simps [meta] nonzero_singleton [meta]},
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   919
      singleton (Morphism.fact phi) @{thm nth_el_Cons [meta]},
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   920
      singleton (Morphism.fact phi) @{thm feval_eq}))))
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   921
\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   922
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   923
end
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   924
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   925
method_setup field = \<open>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   926
  Scan.lift (Args.mode "prems") -- Attrib.thms >> (fn (in_prem, thms) => fn ctxt =>
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   927
    SIMPLE_METHOD' (Field_Tac.field_tac in_prem ctxt THEN' Ring_Tac.ring_tac in_prem thms ctxt))
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   928
\<close> "reduce equations over fields to equations over rings"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   929
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   930
end