src/FOL/ex/Locale_Test/Locale_Test1.thy
author ballarin
Wed, 18 Nov 2015 21:18:33 +0100
changeset 61701 e89cfc004f18
parent 61605 1bf7b186542e
child 62020 5d208fd2507d
permissions -rw-r--r--
Refine the supression of abbreviations for morphisms that are not identities.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
     1
(*  Title:      FOL/ex/Locale_Test/Locale_Test1.thy
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
     2
    Author:     Clemens Ballarin, TU Muenchen
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
     3
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
     4
Test environment for the locale implementation.
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
     5
*)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
     6
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
     7
theory Locale_Test1
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
     8
imports FOL
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
     9
begin
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    10
55380
4de48353034e prefer vacuous definitional type classes over axiomatic ones;
wenzelm
parents: 53367
diff changeset
    11
typedecl int
4de48353034e prefer vacuous definitional type classes over axiomatic ones;
wenzelm
parents: 53367
diff changeset
    12
instance int :: "term" ..
4de48353034e prefer vacuous definitional type classes over axiomatic ones;
wenzelm
parents: 53367
diff changeset
    13
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    14
consts plus :: "int => int => int" (infixl "+" 60)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    15
  zero :: int ("0")
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    16
  minus :: "int => int" ("- _")
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    17
41779
a68f503805ed modernized specifications;
wenzelm
parents: 41435
diff changeset
    18
axiomatization where
a68f503805ed modernized specifications;
wenzelm
parents: 41435
diff changeset
    19
  int_assoc: "(x + y::int) + z = x + (y + z)" and
a68f503805ed modernized specifications;
wenzelm
parents: 41435
diff changeset
    20
  int_zero: "0 + x = x" and
a68f503805ed modernized specifications;
wenzelm
parents: 41435
diff changeset
    21
  int_minus: "(-x) + x = 0" and
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    22
  int_minus2: "-(-x) = x"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    23
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
    24
section \<open>Inference of parameter types\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    25
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    26
locale param1 = fixes p
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    27
print_locale! param1
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    28
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    29
locale param2 = fixes p :: 'b
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    30
print_locale! param2
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    31
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    32
(*
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    33
locale param_top = param2 r for r :: "'b :: {}"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    34
  Fails, cannot generalise parameter.
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    35
*)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    36
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    37
locale param3 = fixes p (infix ".." 50)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    38
print_locale! param3
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    39
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    40
locale param4 = fixes p :: "'a => 'a => 'a" (infix ".." 50)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    41
print_locale! param4
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    42
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    43
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
    44
subsection \<open>Incremental type constraints\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    45
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    46
locale constraint1 =
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    47
  fixes  prod (infixl "**" 65)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    48
  assumes l_id: "x ** y = x"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    49
  assumes assoc: "(x ** y) ** z = x ** (y ** z)"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    50
print_locale! constraint1
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    51
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    52
locale constraint2 =
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    53
  fixes p and q
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    54
  assumes "p = q"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    55
print_locale! constraint2
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    56
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    57
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
    58
section \<open>Inheritance\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    59
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    60
locale semi =
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    61
  fixes prod (infixl "**" 65)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    62
  assumes assoc: "(x ** y) ** z = x ** (y ** z)"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    63
print_locale! semi thm semi_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    64
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    65
locale lgrp = semi +
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    66
  fixes one and inv
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    67
  assumes lone: "one ** x = x"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    68
    and linv: "inv(x) ** x = one"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    69
print_locale! lgrp thm lgrp_def lgrp_axioms_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    70
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    71
locale add_lgrp = semi "op ++" for sum (infixl "++" 60) +
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    72
  fixes zero and neg
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    73
  assumes lzero: "zero ++ x = x"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    74
    and lneg: "neg(x) ++ x = zero"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    75
print_locale! add_lgrp thm add_lgrp_def add_lgrp_axioms_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    76
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    77
locale rev_lgrp = semi "%x y. y ++ x" for sum (infixl "++" 60)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    78
print_locale! rev_lgrp thm rev_lgrp_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    79
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    80
locale hom = f: semi f + g: semi g for f and g
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    81
print_locale! hom thm hom_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    82
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    83
locale perturbation = semi + d: semi "%x y. delta(x) ** delta(y)" for delta
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    84
print_locale! perturbation thm perturbation_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    85
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    86
locale pert_hom = d1: perturbation f d1 + d2: perturbation f d2 for f d1 d2
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    87
print_locale! pert_hom thm pert_hom_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    88
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
    89
text \<open>Alternative expression, obtaining nicer names in @{text "semi f"}.\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    90
locale pert_hom' = semi f + d1: perturbation f d1 + d2: perturbation f d2 for f d1 d2
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    91
print_locale! pert_hom' thm pert_hom'_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    92
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    93
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
    94
section \<open>Syntax declarations\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    95
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    96
locale logic =
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    97
  fixes land (infixl "&&" 55)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    98
    and lnot ("-- _" [60] 60)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
    99
  assumes assoc: "(x && y) && z = x && (y && z)"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   100
    and notnot: "-- (-- x) = x"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   101
begin
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   102
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   103
definition lor (infixl "||" 50) where
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   104
  "x || y = --(-- x && -- y)"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   105
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   106
end
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   107
print_locale! logic
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   108
49756
28e37eab4e6f added some ad-hoc namespace prefixes to avoid duplicate facts;
wenzelm
parents: 43543
diff changeset
   109
locale use_decl = l: logic + semi "op ||"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   110
print_locale! use_decl thm use_decl_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   111
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   112
locale extra_type =
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   113
  fixes a :: 'a
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   114
    and P :: "'a => 'b => o"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   115
begin
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   116
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   117
definition test :: "'a => o"
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   118
  where "test(x) \<longleftrightarrow> (\<forall>b. P(x, b))"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   119
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   120
end
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   121
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   122
term extra_type.test thm extra_type.test_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   123
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   124
interpretation var?: extra_type "0" "%x y. x = 0" .
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   125
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   126
thm var.test_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   127
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   128
61701
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   129
text \<open>Under which circumstances notation remains active.\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   130
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   131
ML \<open>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   132
  fun check_syntax ctxt thm expected =
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   133
    let
56036
6c3fc3b5592a enabled test in PIDE interaction;
wenzelm
parents: 55380
diff changeset
   134
      val obtained =
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 60770
diff changeset
   135
        Print_Mode.setmp [] (Thm.string_of_thm (Config.put show_markup false ctxt)) thm;
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   136
    in
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   137
      if obtained <> expected
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   138
      then error ("Theorem syntax '" ^ obtained ^ "' obtained, but '" ^ expected ^ "' expected.")
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   139
      else ()
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   140
    end;
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   141
\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   142
41305
42967939ea81 actually enable show_hyps option, unlike local_setup in 6da953d30f48 which merely affects the (temporary) auxiliary context;
wenzelm
parents: 41272
diff changeset
   143
declare [[show_hyps]]
41271
6da953d30f48 Enable show_hyps, which appears to be set in batch mode but in an interactive session.
ballarin
parents: 39557
diff changeset
   144
61701
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   145
locale "syntax" =
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   146
  fixes p1 :: "'a => 'b"
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   147
    and p2 :: "'b => o"
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   148
begin
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   149
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   150
definition d1 :: "'a => o" ("D1'(_')") where "d1(x) \<longleftrightarrow> \<not> p2(p1(x))"
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   151
definition d2 :: "'b => o" ("D2'(_')") where "d2(x) \<longleftrightarrow> \<not> p2(x)"
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   152
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   153
thm d1_def d2_def
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   154
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   155
end
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   156
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   157
thm syntax.d1_def syntax.d2_def
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   158
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   159
locale syntax' = "syntax" p1 p2 for p1 :: "'a => 'a" and p2 :: "'a => o"
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   160
begin
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   161
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   162
thm d1_def d2_def  (* should print as "D1(?x) <-> ..." and "D2(?x) <-> ..." *)
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   163
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   164
ML \<open>
61701
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   165
  check_syntax @{context} @{thm d1_def} "D1(?x) \<longleftrightarrow> \<not> p2(p1(?x))";
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   166
  check_syntax @{context} @{thm d2_def} "D2(?x) \<longleftrightarrow> \<not> p2(?x)";
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   167
\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   168
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   169
end
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   170
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   171
locale syntax'' = "syntax" p3 p2 for p3 :: "'a => 'b" and p2 :: "'b => o"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   172
begin
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   173
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   174
thm d1_def d2_def
61701
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   175
  (* should print as "d1(?x) <-> ..." and "D2(?x) <-> ..." *)
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   176
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   177
ML \<open>
61701
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   178
  check_syntax @{context} @{thm d1_def} "d1(?x) \<longleftrightarrow> \<not> p2(p3(?x))";
e89cfc004f18 Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents: 61605
diff changeset
   179
  check_syntax @{context} @{thm d2_def} "D2(?x) \<longleftrightarrow> \<not> p2(?x)";
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   180
\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   181
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   182
end
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   183
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   184
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   185
section \<open>Foundational versions of theorems\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   186
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   187
thm logic.assoc
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   188
thm logic.lor_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   189
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   190
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   191
section \<open>Defines\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   192
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   193
locale logic_def =
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   194
  fixes land (infixl "&&" 55)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   195
    and lor (infixl "||" 50)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   196
    and lnot ("-- _" [60] 60)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   197
  assumes assoc: "(x && y) && z = x && (y && z)"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   198
    and notnot: "-- (-- x) = x"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   199
  defines "x || y == --(-- x && -- y)"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   200
begin
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   201
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   202
thm lor_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   203
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   204
lemma "x || y = --(-- x && --y)"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   205
  by (unfold lor_def) (rule refl)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   206
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   207
end
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   208
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   209
(* Inheritance of defines *)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   210
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   211
locale logic_def2 = logic_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   212
begin
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   213
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   214
lemma "x || y = --(-- x && --y)"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   215
  by (unfold lor_def) (rule refl)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   216
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   217
end
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   218
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   219
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   220
section \<open>Notes\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   221
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   222
(* A somewhat arcane homomorphism example *)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   223
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   224
definition semi_hom where
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   225
  "semi_hom(prod, sum, h) \<longleftrightarrow> (\<forall>x y. h(prod(x, y)) = sum(h(x), h(y)))"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   226
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   227
lemma semi_hom_mult:
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   228
  "semi_hom(prod, sum, h) \<Longrightarrow> h(prod(x, y)) = sum(h(x), h(y))"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   229
  by (simp add: semi_hom_def)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   230
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   231
locale semi_hom_loc = prod: semi prod + sum: semi sum
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   232
  for prod and sum and h +
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   233
  assumes semi_homh: "semi_hom(prod, sum, h)"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   234
  notes semi_hom_mult = semi_hom_mult [OF semi_homh]
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   235
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   236
thm semi_hom_loc.semi_hom_mult
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   237
(* unspecified, attribute not applied in backgroud theory !!! *)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   238
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   239
lemma (in semi_hom_loc) "h(prod(x, y)) = sum(h(x), h(y))"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   240
  by (rule semi_hom_mult)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   241
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   242
(* Referring to facts from within a context specification *)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   243
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   244
lemma
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   245
  assumes x: "P \<longleftrightarrow> P"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   246
  notes y = x
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   247
  shows True ..
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   248
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   249
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   250
section \<open>Theorem statements\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   251
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   252
lemma (in lgrp) lcancel:
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   253
  "x ** y = x ** z \<longleftrightarrow> y = z"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   254
proof
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   255
  assume "x ** y = x ** z"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   256
  then have "inv(x) ** x ** y = inv(x) ** x ** z" by (simp add: assoc)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   257
  then show "y = z" by (simp add: lone linv)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   258
qed simp
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   259
print_locale! lgrp
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   260
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   261
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   262
locale rgrp = semi +
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   263
  fixes one and inv
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   264
  assumes rone: "x ** one = x"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   265
    and rinv: "x ** inv(x) = one"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   266
begin
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   267
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   268
lemma rcancel:
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   269
  "y ** x = z ** x \<longleftrightarrow> y = z"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   270
proof
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   271
  assume "y ** x = z ** x"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   272
  then have "y ** (x ** inv(x)) = z ** (x ** inv(x))"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   273
    by (simp add: assoc [symmetric])
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   274
  then show "y = z" by (simp add: rone rinv)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   275
qed simp
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   276
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   277
end
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   278
print_locale! rgrp
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   279
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   280
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   281
subsection \<open>Patterns\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   282
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   283
lemma (in rgrp)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   284
  assumes "y ** x = z ** x" (is ?a)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   285
  shows "y = z" (is ?t)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   286
proof -
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   287
  txt \<open>Weird proof involving patterns from context element and conclusion.\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   288
  {
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   289
    assume ?a
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   290
    then have "y ** (x ** inv(x)) = z ** (x ** inv(x))"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   291
      by (simp add: assoc [symmetric])
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   292
    then have ?t by (simp add: rone rinv)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   293
  }
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   294
  note x = this
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   295
  show ?t by (rule x [OF \<open>?a\<close>])
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   296
qed
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   297
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   298
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   299
section \<open>Interpretation between locales: sublocales\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   300
61565
352c73a689da Qualifiers in locale expressions default to mandatory regardless of the command.
ballarin
parents: 61489
diff changeset
   301
sublocale lgrp < right?: rgrp
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   302
print_facts
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   303
proof unfold_locales
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   304
  {
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   305
    fix x
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   306
    have "inv(x) ** x ** one = inv(x) ** x" by (simp add: linv lone)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   307
    then show "x ** one = x" by (simp add: assoc lcancel)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   308
  }
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   309
  note rone = this
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   310
  {
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   311
    fix x
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   312
    have "inv(x) ** x ** inv(x) = inv(x) ** one"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   313
      by (simp add: linv lone rone)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   314
    then show "x ** inv(x) = one" by (simp add: assoc lcancel)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   315
  }
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   316
qed
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   317
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   318
(* effect on printed locale *)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   319
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   320
print_locale! lgrp
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   321
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   322
(* use of derived theorem *)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   323
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   324
lemma (in lgrp)
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   325
  "y ** x = z ** x \<longleftrightarrow> y = z"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   326
  apply (rule rcancel)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   327
  done
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   328
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   329
(* circular interpretation *)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   330
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   331
sublocale rgrp < left: lgrp
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   332
proof unfold_locales
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   333
  {
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   334
    fix x
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   335
    have "one ** (x ** inv(x)) = x ** inv(x)" by (simp add: rinv rone)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   336
    then show "one ** x = x" by (simp add: assoc [symmetric] rcancel)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   337
  }
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   338
  note lone = this
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   339
  {
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   340
    fix x
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   341
    have "inv(x) ** (x ** inv(x)) = one ** inv(x)"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   342
      by (simp add: rinv lone rone)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   343
    then show "inv(x) ** x = one" by (simp add: assoc [symmetric] rcancel)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   344
  }
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   345
qed
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   346
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   347
(* effect on printed locale *)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   348
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   349
print_locale! rgrp
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   350
print_locale! lgrp
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   351
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   352
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   353
(* Duality *)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   354
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   355
locale order =
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   356
  fixes less :: "'a => 'a => o" (infix "<<" 50)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   357
  assumes refl: "x << x"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   358
    and trans: "[| x << y; y << z |] ==> x << z"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   359
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   360
sublocale order < dual: order "%x y. y << x"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   361
  apply unfold_locales apply (rule refl) apply (blast intro: trans)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   362
  done
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   363
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   364
print_locale! order  (* Only two instances of order. *)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   365
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   366
locale order' =
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   367
  fixes less :: "'a => 'a => o" (infix "<<" 50)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   368
  assumes refl: "x << x"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   369
    and trans: "[| x << y; y << z |] ==> x << z"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   370
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   371
locale order_with_def = order'
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   372
begin
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   373
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   374
definition greater :: "'a => 'a => o" (infix ">>" 50) where
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   375
  "x >> y \<longleftrightarrow> y << x"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   376
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   377
end
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   378
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   379
sublocale order_with_def < dual: order' "op >>"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   380
  apply unfold_locales
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   381
  unfolding greater_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   382
  apply (rule refl) apply (blast intro: trans)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   383
  done
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   384
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   385
print_locale! order_with_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   386
(* Note that decls come after theorems that make use of them. *)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   387
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   388
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   389
(* locale with many parameters ---
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   390
   interpretations generate alternating group A5 *)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   391
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   392
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   393
locale A5 =
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   394
  fixes A and B and C and D and E
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   395
  assumes eq: "A \<longleftrightarrow> B \<longleftrightarrow> C \<longleftrightarrow> D \<longleftrightarrow> E"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   396
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   397
sublocale A5 < 1: A5 _ _ D E C
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   398
print_facts
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   399
  using eq apply (blast intro: A5.intro) done
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   400
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   401
sublocale A5 < 2: A5 C _ E _ A
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   402
print_facts
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   403
  using eq apply (blast intro: A5.intro) done
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   404
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   405
sublocale A5 < 3: A5 B C A _ _
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   406
print_facts
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   407
  using eq apply (blast intro: A5.intro) done
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   408
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   409
(* Any even permutation of parameters is subsumed by the above. *)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   410
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   411
print_locale! A5
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   412
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   413
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   414
(* Free arguments of instance *)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   415
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   416
locale trivial =
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   417
  fixes P and Q :: o
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   418
  assumes Q: "P \<longleftrightarrow> P \<longleftrightarrow> Q"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   419
begin
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   420
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   421
lemma Q_triv: "Q" using Q by fast
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   422
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   423
end
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   424
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   425
sublocale trivial < x: trivial x _
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   426
  apply unfold_locales using Q by fast
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   427
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   428
print_locale! trivial
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   429
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   430
context trivial begin thm x.Q [where ?x = True] end
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   431
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   432
sublocale trivial < y: trivial Q Q
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   433
  by unfold_locales
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   434
  (* Succeeds since previous interpretation is more general. *)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   435
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   436
print_locale! trivial  (* No instance for y created (subsumed). *)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   437
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   438
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   439
subsection \<open>Sublocale, then interpretation in theory\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   440
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   441
interpretation int?: lgrp "op +" "0" "minus"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   442
proof unfold_locales
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   443
qed (rule int_assoc int_zero int_minus)+
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   444
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   445
thm int.assoc int.semi_axioms
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   446
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   447
interpretation int2?: semi "op +"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   448
  by unfold_locales  (* subsumed, thm int2.assoc not generated *)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   449
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   450
ML \<open>(Global_Theory.get_thms @{theory} "int2.assoc";
56138
f42de6d8ed8e do not test details of error messages;
wenzelm
parents: 56036
diff changeset
   451
    raise Fail "thm int2.assoc was generated")
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   452
  handle ERROR _ => ([]:thm list);\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   453
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   454
thm int.lone int.right.rone
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   455
  (* the latter comes through the sublocale relation *)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   456
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   457
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   458
subsection \<open>Interpretation in theory, then sublocale\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   459
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   460
interpretation fol: logic "op +" "minus"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   461
  by unfold_locales (rule int_assoc int_minus2)+
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   462
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   463
locale logic2 =
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   464
  fixes land (infixl "&&" 55)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   465
    and lnot ("-- _" [60] 60)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   466
  assumes assoc: "(x && y) && z = x && (y && z)"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   467
    and notnot: "-- (-- x) = x"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   468
begin
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   469
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   470
definition lor (infixl "||" 50) where
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   471
  "x || y = --(-- x && -- y)"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   472
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   473
end
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   474
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   475
sublocale logic < two: logic2
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   476
  by unfold_locales (rule assoc notnot)+
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   477
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   478
thm fol.two.assoc
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   479
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   480
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   481
subsection \<open>Declarations and sublocale\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   482
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   483
locale logic_a = logic
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   484
locale logic_b = logic
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   485
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   486
sublocale logic_a < logic_b
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   487
  by unfold_locales
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   488
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   489
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   490
subsection \<open>Interpretation\<close>
53366
c8c548d83862 Terminology: mixin -> rewrite morphism.
ballarin
parents: 51515
diff changeset
   491
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   492
subsection \<open>Rewrite morphism\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   493
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   494
locale logic_o =
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   495
  fixes land (infixl "&&" 55)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   496
    and lnot ("-- _" [60] 60)
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   497
  assumes assoc_o: "(x && y) && z \<longleftrightarrow> x && (y && z)"
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   498
    and notnot_o: "-- (-- x) \<longleftrightarrow> x"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   499
begin
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   500
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   501
definition lor_o (infixl "||" 50) where
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   502
  "x || y \<longleftrightarrow> --(-- x && -- y)"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   503
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   504
end
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   505
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   506
interpretation x: logic_o "op \<and>" "Not"
61566
c3d6e570ccef Keyword 'rewrites' identifies rewrite morphisms.
ballarin
parents: 61565
diff changeset
   507
  rewrites bool_logic_o: "x.lor_o(x, y) \<longleftrightarrow> x \<or> y"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   508
proof -
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   509
  show bool_logic_o: "PROP logic_o(op \<and>, Not)" by unfold_locales fast+
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   510
  show "logic_o.lor_o(op \<and>, Not, x, y) \<longleftrightarrow> x \<or> y"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   511
    by (unfold logic_o.lor_o_def [OF bool_logic_o]) fast
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   512
qed
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   513
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   514
thm x.lor_o_def bool_logic_o
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   515
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   516
lemma lor_triv: "z \<longleftrightarrow> z" ..
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   517
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   518
lemma (in logic_o) lor_triv: "x || y \<longleftrightarrow> x || y" by fast
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   519
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   520
thm lor_triv [where z = True] (* Check strict prefix. *)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   521
  x.lor_triv
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   522
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   523
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   524
subsection \<open>Inheritance of rewrite morphisms\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   525
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   526
locale reflexive =
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   527
  fixes le :: "'a => 'a => o" (infix "\<sqsubseteq>" 50)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   528
  assumes refl: "x \<sqsubseteq> x"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   529
begin
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   530
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   531
definition less (infix "\<sqsubset>" 50) where "x \<sqsubset> y \<longleftrightarrow> x \<sqsubseteq> y \<and> x \<noteq> y"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   532
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   533
end
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   534
41779
a68f503805ed modernized specifications;
wenzelm
parents: 41435
diff changeset
   535
axiomatization
a68f503805ed modernized specifications;
wenzelm
parents: 41435
diff changeset
   536
  gle :: "'a => 'a => o" and gless :: "'a => 'a => o" and
a68f503805ed modernized specifications;
wenzelm
parents: 41435
diff changeset
   537
  gle' :: "'a => 'a => o" and gless' :: "'a => 'a => o"
a68f503805ed modernized specifications;
wenzelm
parents: 41435
diff changeset
   538
where
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   539
  grefl: "gle(x, x)" and gless_def: "gless(x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y" and
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   540
  grefl': "gle'(x, x)" and gless'_def: "gless'(x, y) \<longleftrightarrow> gle'(x, y) \<and> x \<noteq> y"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   541
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   542
text \<open>Setup\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   543
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   544
locale mixin = reflexive
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   545
begin
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   546
lemmas less_thm = less_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   547
end
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   548
61566
c3d6e570ccef Keyword 'rewrites' identifies rewrite morphisms.
ballarin
parents: 61565
diff changeset
   549
interpretation le: mixin gle rewrites "reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   550
proof -
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   551
  show "mixin(gle)" by unfold_locales (rule grefl)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   552
  note reflexive = this[unfolded mixin_def]
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   553
  show "reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   554
    by (simp add: reflexive.less_def[OF reflexive] gless_def)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   555
qed
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   556
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   557
text \<open>Rewrite morphism propagated along the locale hierarchy\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   558
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   559
locale mixin2 = mixin
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   560
begin
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   561
lemmas less_thm2 = less_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   562
end
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   563
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   564
interpretation le: mixin2 gle
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   565
  by unfold_locales
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   566
53366
c8c548d83862 Terminology: mixin -> rewrite morphism.
ballarin
parents: 51515
diff changeset
   567
thm le.less_thm2  (* rewrite morphism applied *)
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   568
lemma "gless(x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   569
  by (rule le.less_thm2)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   570
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   571
text \<open>Rewrite morphism does not leak to a side branch.\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   572
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   573
locale mixin3 = reflexive
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   574
begin
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   575
lemmas less_thm3 = less_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   576
end
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   577
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   578
interpretation le: mixin3 gle
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   579
  by unfold_locales
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   580
53366
c8c548d83862 Terminology: mixin -> rewrite morphism.
ballarin
parents: 51515
diff changeset
   581
thm le.less_thm3  (* rewrite morphism not applied *)
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   582
lemma "reflexive.less(gle, x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y" by (rule le.less_thm3)
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   583
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   584
text \<open>Rewrite morphism only available in original context\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   585
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   586
locale mixin4_base = reflexive
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   587
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   588
locale mixin4_mixin = mixin4_base
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   589
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   590
interpretation le: mixin4_mixin gle
61566
c3d6e570ccef Keyword 'rewrites' identifies rewrite morphisms.
ballarin
parents: 61565
diff changeset
   591
  rewrites "reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   592
proof -
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   593
  show "mixin4_mixin(gle)" by unfold_locales (rule grefl)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   594
  note reflexive = this[unfolded mixin4_mixin_def mixin4_base_def mixin_def]
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   595
  show "reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   596
    by (simp add: reflexive.less_def[OF reflexive] gless_def)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   597
qed
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   598
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   599
locale mixin4_copy = mixin4_base
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   600
begin
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   601
lemmas less_thm4 = less_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   602
end
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   603
61565
352c73a689da Qualifiers in locale expressions default to mandatory regardless of the command.
ballarin
parents: 61489
diff changeset
   604
locale mixin4_combined = le1?: mixin4_mixin le' + le2?: mixin4_copy le for le' le
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   605
begin
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   606
lemmas less_thm4' = less_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   607
end
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   608
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   609
interpretation le4: mixin4_combined gle' gle
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   610
  by unfold_locales (rule grefl')
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   611
53366
c8c548d83862 Terminology: mixin -> rewrite morphism.
ballarin
parents: 51515
diff changeset
   612
thm le4.less_thm4' (* rewrite morphism not applied *)
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   613
lemma "reflexive.less(gle, x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   614
  by (rule le4.less_thm4')
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   615
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   616
text \<open>Inherited rewrite morphism applied to new theorem\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   617
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   618
locale mixin5_base = reflexive
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   619
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   620
locale mixin5_inherited = mixin5_base
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   621
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   622
interpretation le5: mixin5_base gle
61566
c3d6e570ccef Keyword 'rewrites' identifies rewrite morphisms.
ballarin
parents: 61565
diff changeset
   623
  rewrites "reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   624
proof -
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   625
  show "mixin5_base(gle)" by unfold_locales
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   626
  note reflexive = this[unfolded mixin5_base_def mixin_def]
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   627
  show "reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   628
    by (simp add: reflexive.less_def[OF reflexive] gless_def)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   629
qed
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   630
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   631
interpretation le5: mixin5_inherited gle
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   632
  by unfold_locales
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   633
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   634
lemmas (in mixin5_inherited) less_thm5 = less_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   635
53366
c8c548d83862 Terminology: mixin -> rewrite morphism.
ballarin
parents: 51515
diff changeset
   636
thm le5.less_thm5  (* rewrite morphism applied *)
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   637
lemma "gless(x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   638
  by (rule le5.less_thm5)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   639
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   640
text \<open>Rewrite morphism pushed down to existing inherited locale\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   641
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   642
locale mixin6_base = reflexive
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   643
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   644
locale mixin6_inherited = mixin5_base
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   645
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   646
interpretation le6: mixin6_base gle
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   647
  by unfold_locales
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   648
interpretation le6: mixin6_inherited gle
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   649
  by unfold_locales
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   650
interpretation le6: mixin6_base gle
61566
c3d6e570ccef Keyword 'rewrites' identifies rewrite morphisms.
ballarin
parents: 61565
diff changeset
   651
  rewrites "reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   652
proof -
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   653
  show "mixin6_base(gle)" by unfold_locales
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   654
  note reflexive = this[unfolded mixin6_base_def mixin_def]
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   655
  show "reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   656
    by (simp add: reflexive.less_def[OF reflexive] gless_def)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   657
qed
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   658
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   659
lemmas (in mixin6_inherited) less_thm6 = less_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   660
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   661
thm le6.less_thm6  (* mixin applied *)
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   662
lemma "gless(x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   663
  by (rule le6.less_thm6)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   664
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   665
text \<open>Existing rewrite morphism inherited through sublocale relation\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   666
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   667
locale mixin7_base = reflexive
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   668
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   669
locale mixin7_inherited = reflexive
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   670
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   671
interpretation le7: mixin7_base gle
61566
c3d6e570ccef Keyword 'rewrites' identifies rewrite morphisms.
ballarin
parents: 61565
diff changeset
   672
  rewrites "reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   673
proof -
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   674
  show "mixin7_base(gle)" by unfold_locales
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   675
  note reflexive = this[unfolded mixin7_base_def mixin_def]
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   676
  show "reflexive.less(gle, x, y) \<longleftrightarrow> gless(x, y)"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   677
    by (simp add: reflexive.less_def[OF reflexive] gless_def)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   678
qed
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   679
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   680
interpretation le7: mixin7_inherited gle
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   681
  by unfold_locales
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   682
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   683
lemmas (in mixin7_inherited) less_thm7 = less_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   684
53366
c8c548d83862 Terminology: mixin -> rewrite morphism.
ballarin
parents: 51515
diff changeset
   685
thm le7.less_thm7  (* before, rewrite morphism not applied *)
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   686
lemma "reflexive.less(gle, x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   687
  by (rule le7.less_thm7)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   688
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   689
sublocale mixin7_inherited < mixin7_base
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   690
  by unfold_locales
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   691
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   692
lemmas (in mixin7_inherited) less_thm7b = less_def
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   693
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   694
thm le7.less_thm7b  (* after, mixin applied *)
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   695
lemma "gless(x, y) \<longleftrightarrow> gle(x, y) \<and> x \<noteq> y"
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   696
  by (rule le7.less_thm7b)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   697
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   698
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   699
text \<open>This locale will be interpreted in later theories.\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   700
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   701
locale mixin_thy_merge = le: reflexive le + le': reflexive le' for le le'
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   702
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   703
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   704
subsection \<open>Rewrite morphisms in sublocale\<close>
41272
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   705
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   706
text \<open>Simulate a specification of left groups where unit and inverse are defined
41272
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   707
  rather than specified.  This is possible, but not in FOL, due to the lack of a
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   708
  selection operator.\<close>
41272
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   709
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   710
axiomatization glob_one and glob_inv
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   711
  where glob_lone: "prod(glob_one(prod), x) = x"
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   712
    and glob_linv: "prod(glob_inv(prod, x), x) = glob_one(prod)"
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   713
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   714
locale dgrp = semi
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   715
begin
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   716
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   717
definition one where "one = glob_one(prod)"
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   718
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   719
lemma lone: "one ** x = x"
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   720
  unfolding one_def by (rule glob_lone)
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   721
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   722
definition inv where "inv(x) = glob_inv(prod, x)"
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   723
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   724
lemma linv: "inv(x) ** x = one"
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   725
  unfolding one_def inv_def by (rule glob_linv)
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   726
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   727
end
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   728
61565
352c73a689da Qualifiers in locale expressions default to mandatory regardless of the command.
ballarin
parents: 61489
diff changeset
   729
sublocale lgrp < "def"?: dgrp
61566
c3d6e570ccef Keyword 'rewrites' identifies rewrite morphisms.
ballarin
parents: 61565
diff changeset
   730
  rewrites one_equation: "dgrp.one(prod) = one" and inv_equation: "dgrp.inv(prod, x) = inv(x)"
41272
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   731
proof -
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   732
  show "dgrp(prod)" by unfold_locales
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   733
  from this interpret d: dgrp .
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   734
  -- Unit
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   735
  have "dgrp.one(prod) = glob_one(prod)" by (rule d.one_def)
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   736
  also have "... = glob_one(prod) ** one" by (simp add: rone)
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   737
  also have "... = one" by (simp add: glob_lone)
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   738
  finally show "dgrp.one(prod) = one" .
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   739
  -- Inverse
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   740
  then have "dgrp.inv(prod, x) ** x = inv(x) ** x" by (simp add: glob_linv d.linv linv)
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   741
  then show "dgrp.inv(prod, x) = inv(x)" by (simp add: rcancel)
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   742
qed
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   743
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   744
print_locale! lgrp
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   745
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   746
context lgrp begin
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   747
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   748
text \<open>Equations stored in target\<close>
41272
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   749
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   750
lemma "dgrp.one(prod) = one" by (rule one_equation)
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   751
lemma "dgrp.inv(prod, x) = inv(x)" by (rule inv_equation)
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   752
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   753
text \<open>Rewrite morphisms applied\<close>
41272
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   754
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   755
lemma "one = glob_one(prod)" by (rule one_def)
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   756
lemma "inv(x) = glob_inv(prod, x)" by (rule inv_def)
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   757
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   758
end
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   759
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   760
text \<open>Interpreted versions\<close>
41272
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   761
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   762
lemma "0 = glob_one (op +)" by (rule int.def.one_def)
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   763
lemma "- x = glob_inv(op +, x)" by (rule int.def.inv_def)
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   764
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   765
text \<open>Roundup applies rewrite morphisms at declaration level in DFS tree\<close>
51515
c3eb0b517ced Fix issue related to mixins in roundup.
ballarin
parents: 49756
diff changeset
   766
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   767
locale roundup = fixes x assumes true: "x \<longleftrightarrow> True"
51515
c3eb0b517ced Fix issue related to mixins in roundup.
ballarin
parents: 49756
diff changeset
   768
61566
c3d6e570ccef Keyword 'rewrites' identifies rewrite morphisms.
ballarin
parents: 61565
diff changeset
   769
sublocale roundup \<subseteq> sub: roundup x rewrites "x \<longleftrightarrow> True \<and> True"
51515
c3eb0b517ced Fix issue related to mixins in roundup.
ballarin
parents: 49756
diff changeset
   770
  apply unfold_locales apply (simp add: true) done
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   771
lemma (in roundup) "True \<and> True \<longleftrightarrow> True" by (rule sub.true)
51515
c3eb0b517ced Fix issue related to mixins in roundup.
ballarin
parents: 49756
diff changeset
   772
41272
b806a7678083 Add mixins to locale dependencies.
ballarin
parents: 41271
diff changeset
   773
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   774
section \<open>Interpretation in named contexts\<close>
53367
1f383542226b New test case: interpretation in named contexts is not persistent.
ballarin
parents: 53366
diff changeset
   775
1f383542226b New test case: interpretation in named contexts is not persistent.
ballarin
parents: 53366
diff changeset
   776
locale container
1f383542226b New test case: interpretation in named contexts is not persistent.
ballarin
parents: 53366
diff changeset
   777
begin
61605
1bf7b186542e qualifier is mandatory by default;
wenzelm
parents: 61566
diff changeset
   778
interpretation "private": roundup True by unfold_locales rule
53367
1f383542226b New test case: interpretation in named contexts is not persistent.
ballarin
parents: 53366
diff changeset
   779
lemmas true_copy = private.true
1f383542226b New test case: interpretation in named contexts is not persistent.
ballarin
parents: 53366
diff changeset
   780
end
1f383542226b New test case: interpretation in named contexts is not persistent.
ballarin
parents: 53366
diff changeset
   781
1f383542226b New test case: interpretation in named contexts is not persistent.
ballarin
parents: 53366
diff changeset
   782
context container begin
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   783
ML \<open>(Context.>> (fn generic => let val context = Context.proof_of generic
53367
1f383542226b New test case: interpretation in named contexts is not persistent.
ballarin
parents: 53366
diff changeset
   784
  val _ = Proof_Context.get_thms context "private.true" in generic end);
56138
f42de6d8ed8e do not test details of error messages;
wenzelm
parents: 56036
diff changeset
   785
  raise Fail "thm private.true was persisted")
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   786
  handle ERROR _ => ([]:thm list);\<close>
53367
1f383542226b New test case: interpretation in named contexts is not persistent.
ballarin
parents: 53366
diff changeset
   787
thm true_copy
1f383542226b New test case: interpretation in named contexts is not persistent.
ballarin
parents: 53366
diff changeset
   788
end
1f383542226b New test case: interpretation in named contexts is not persistent.
ballarin
parents: 53366
diff changeset
   789
1f383542226b New test case: interpretation in named contexts is not persistent.
ballarin
parents: 53366
diff changeset
   790
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59897
diff changeset
   791
section \<open>Interpretation in proofs\<close>
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   792
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   793
lemma True
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   794
proof
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   795
  interpret "local": lgrp "op +" "0" "minus"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   796
    by unfold_locales  (* subsumed *)
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   797
  {
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   798
    fix zero :: int
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   799
    assume "!!x. zero + x = x" "!!x. (-x) + x = zero"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   800
    then interpret local_fixed: lgrp "op +" zero "minus"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   801
      by unfold_locales
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   802
    thm local_fixed.lone
41435
12585dfb86fe Diagnostic command to show locale dependencies.
ballarin
parents: 41305
diff changeset
   803
    print_dependencies! lgrp "op +" 0 minus + lgrp "op +" zero minus
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   804
  }
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   805
  assume "!!x zero. zero + x = x" "!!x zero. (-x) + x = zero"
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   806
  then interpret local_free: lgrp "op +" zero "minus" for zero
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   807
    by unfold_locales
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   808
  thm local_free.lone [where ?zero = 0]
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   809
qed
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   810
38108
b4115423c049 Interpretation in proofs supports mixins.
ballarin
parents: 37146
diff changeset
   811
lemma True
b4115423c049 Interpretation in proofs supports mixins.
ballarin
parents: 37146
diff changeset
   812
proof
b4115423c049 Interpretation in proofs supports mixins.
ballarin
parents: 37146
diff changeset
   813
  {
b4115423c049 Interpretation in proofs supports mixins.
ballarin
parents: 37146
diff changeset
   814
    fix pand and pnot and por
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   815
    assume passoc: "\<And>x y z. pand(pand(x, y), z) \<longleftrightarrow> pand(x, pand(y, z))"
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   816
      and pnotnot: "\<And>x. pnot(pnot(x)) \<longleftrightarrow> x"
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   817
      and por_def: "\<And>x y. por(x, y) \<longleftrightarrow> pnot(pand(pnot(x), pnot(y)))"
38108
b4115423c049 Interpretation in proofs supports mixins.
ballarin
parents: 37146
diff changeset
   818
    interpret loc: logic_o pand pnot
61566
c3d6e570ccef Keyword 'rewrites' identifies rewrite morphisms.
ballarin
parents: 61565
diff changeset
   819
      rewrites por_eq: "\<And>x y. logic_o.lor_o(pand, pnot, x, y) \<longleftrightarrow> por(x, y)"  (* FIXME *)
38108
b4115423c049 Interpretation in proofs supports mixins.
ballarin
parents: 37146
diff changeset
   820
    proof -
b4115423c049 Interpretation in proofs supports mixins.
ballarin
parents: 37146
diff changeset
   821
      show logic_o: "PROP logic_o(pand, pnot)" using passoc pnotnot by unfold_locales
b4115423c049 Interpretation in proofs supports mixins.
ballarin
parents: 37146
diff changeset
   822
      fix x y
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   823
      show "logic_o.lor_o(pand, pnot, x, y) \<longleftrightarrow> por(x, y)"
38108
b4115423c049 Interpretation in proofs supports mixins.
ballarin
parents: 37146
diff changeset
   824
        by (unfold logic_o.lor_o_def [OF logic_o]) (rule por_def [symmetric])
b4115423c049 Interpretation in proofs supports mixins.
ballarin
parents: 37146
diff changeset
   825
    qed
38109
06fd1914b902 print_interps shows interpretations in proofs.
ballarin
parents: 38108
diff changeset
   826
    print_interps logic_o
61489
b8d375aee0df more symbols;
wenzelm
parents: 61487
diff changeset
   827
    have "\<And>x y. por(x, y) \<longleftrightarrow> pnot(pand(pnot(x), pnot(y)))" by (rule loc.lor_o_def)
38108
b4115423c049 Interpretation in proofs supports mixins.
ballarin
parents: 37146
diff changeset
   828
  }
b4115423c049 Interpretation in proofs supports mixins.
ballarin
parents: 37146
diff changeset
   829
qed
b4115423c049 Interpretation in proofs supports mixins.
ballarin
parents: 37146
diff changeset
   830
37134
29bd6c2ffba8 Revise locale test theory layout.
ballarin
parents:
diff changeset
   831
end