| author | haftmann | 
| Tue, 07 Oct 2008 16:07:33 +0200 | |
| changeset 28522 | eacb54d9e78d | 
| parent 28235 | 89e4d2232ed2 | 
| child 28611 | 983c1855a7af | 
| permissions | -rw-r--r-- | 
| 15596 | 1 | (* Title: FOL/ex/LocaleTest.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Clemens Ballarin | |
| 4 | Copyright (c) 2005 by Clemens Ballarin | |
| 5 | ||
| 6 | Collection of regression tests for locales. | |
| 7 | *) | |
| 8 | ||
| 16102 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
15837diff
changeset | 9 | header {* Test of Locale Interpretation *}
 | 
| 15596 | 10 | |
| 16102 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
15837diff
changeset | 11 | theory LocaleTest | 
| 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
15837diff
changeset | 12 | imports FOL | 
| 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
15837diff
changeset | 13 | begin | 
| 15596 | 14 | |
| 16168 
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
 ballarin parents: 
16102diff
changeset | 15 | ML {* set Toplevel.debug *}
 | 
| 15696 | 16 | ML {* set show_hyps *}
 | 
| 17 | ML {* set show_sorts *}
 | |
| 18 | ||
| 17139 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 19 | ML {*
 | 
| 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 20 | fun check_thm name = let | 
| 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 21 | val thy = the_context (); | 
| 26343 
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
 wenzelm parents: 
26336diff
changeset | 22 | val thm = PureThy.get_thm thy name; | 
| 17139 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 23 |     val {prop, hyps, ...} = rep_thm thm;
 | 
| 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 24 | val prems = Logic.strip_imp_prems prop; | 
| 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 25 | val _ = if null hyps then () | 
| 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 26 |         else error ("Theorem " ^ quote name ^ " has meta hyps.\n" ^
 | 
| 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 27 | "Consistency check of locales package failed."); | 
| 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 28 | val _ = if null prems then () | 
| 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 29 |         else error ("Theorem " ^ quote name ^ " has premises.\n" ^
 | 
| 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 30 | "Consistency check of locales package failed."); | 
| 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 31 | in () end; | 
| 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 32 | *} | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 33 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 34 | section {* Context Elements and Locale Expressions *}
 | 
| 16102 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
15837diff
changeset | 35 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 36 | text {* Naming convention for global objects: prefixes L and l *}
 | 
| 16168 
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
 ballarin parents: 
16102diff
changeset | 37 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 38 | subsection {* Renaming with Syntax *}
 | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 39 | |
| 27681 | 40 | locale LS = var mult + | 
| 16102 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
15837diff
changeset | 41 | assumes "mult(x, y) = mult(y, x)" | 
| 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
15837diff
changeset | 42 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 43 | print_locale LS | 
| 16168 
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
 ballarin parents: 
16102diff
changeset | 44 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 45 | locale LS' = LS mult (infixl "**" 60) | 
| 16102 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
15837diff
changeset | 46 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 47 | print_locale LS' | 
| 16102 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
15837diff
changeset | 48 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 49 | locale LT = var mult (infixl "**" 60) + | 
| 16102 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
15837diff
changeset | 50 | assumes "x ** y = y ** x" | 
| 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
15837diff
changeset | 51 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 52 | locale LU = LT mult (infixl "**" 60) + LT add (infixl "++" 55) + var h + | 
| 16102 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
15837diff
changeset | 53 | assumes hom: "h(x ** y) = h(x) ++ h(y)" | 
| 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
15837diff
changeset | 54 | |
| 25282 | 55 | |
| 56 | (* FIXME: graceful handling of type errors? | |
| 19783 | 57 | locale LY = LT mult (infixl "**" 60) + LT add (binder "++" 55) + var h + | 
| 58 | assumes "mult(x) == add" | |
| 59 | *) | |
| 60 | ||
| 25282 | 61 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 62 | locale LV = LU _ add | 
| 16102 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
15837diff
changeset | 63 | |
| 19783 | 64 | locale LW = LU _ mult (infixl "**" 60) | 
| 65 | ||
| 16168 
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
 ballarin parents: 
16102diff
changeset | 66 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 67 | subsection {* Constrains *}
 | 
| 16102 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
15837diff
changeset | 68 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 69 | locale LZ = fixes a (structure) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 70 | locale LZ' = LZ + | 
| 16168 
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
 ballarin parents: 
16102diff
changeset | 71 | constrains a :: "'a => 'b" | 
| 
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
 ballarin parents: 
16102diff
changeset | 72 | assumes "a (x :: 'a) = a (y)" | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 73 | print_locale LZ' | 
| 16102 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
15837diff
changeset | 74 | |
| 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
15837diff
changeset | 75 | |
| 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
15837diff
changeset | 76 | section {* Interpretation *}
 | 
| 15696 | 77 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 78 | text {* Naming convention for global objects: prefixes I and i *}
 | 
| 15596 | 79 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 80 | text {* interpretation input syntax *}
 | 
| 15596 | 81 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 82 | locale IL | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 83 | locale IM = fixes a and b and c | 
| 15596 | 84 | |
| 28085 
914183e229e9
Interpretation commands no longer accept interpretation attributes.
 ballarin parents: 
27761diff
changeset | 85 | interpretation test: IL + IM a b c [x y z] . | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 86 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 87 | print_interps IL (* output: test *) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 88 | print_interps IM (* output: test *) | 
| 15596 | 89 | |
| 28085 
914183e229e9
Interpretation commands no longer accept interpretation attributes.
 ballarin parents: 
27761diff
changeset | 90 | interpretation test: IL print_interps IM . | 
| 15596 | 91 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 92 | interpretation IL . | 
| 15596 | 93 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 94 | text {* Processing of locale expression *}
 | 
| 15596 | 95 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 96 | locale IA = fixes a assumes asm_A: "a = a" | 
| 15596 | 97 | |
| 27681 | 98 | locale IB = fixes b assumes asm_B [simp]: "b = b" | 
| 15596 | 99 | |
| 28235 
89e4d2232ed2
No interpretation of locale with dangling type frees.
 ballarin parents: 
28085diff
changeset | 100 | locale IC = IA + IB + assumes asm_C: "b = b" | 
| 
89e4d2232ed2
No interpretation of locale with dangling type frees.
 ballarin parents: 
28085diff
changeset | 101 | |
| 
89e4d2232ed2
No interpretation of locale with dangling type frees.
 ballarin parents: 
28085diff
changeset | 102 | locale IC' = IA + IB + assumes asm_C: "c = c" | 
| 
89e4d2232ed2
No interpretation of locale with dangling type frees.
 ballarin parents: 
28085diff
changeset | 103 | (* independent type var in c *) | 
| 15596 | 104 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 105 | locale ID = IA + IB + fixes d defines def_D: "d == (a = b)" | 
| 15596 | 106 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 107 | theorem (in ID) True .. | 
| 15596 | 108 | |
| 109 | typedecl i | |
| 110 | arities i :: "term" | |
| 111 | ||
| 15598 
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
 ballarin parents: 
15596diff
changeset | 112 | |
| 19984 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 ballarin parents: 
19931diff
changeset | 113 | interpretation i1: IC ["X::i" "Y::i"] by unfold_locales auto | 
| 15596 | 114 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 115 | print_interps IA (* output: i1 *) | 
| 15596 | 116 | |
| 15696 | 117 | (* possible accesses *) | 
| 26645 
e114be97befe
Changed naming scheme for theorems generated by interpretations.
 ballarin parents: 
26343diff
changeset | 118 | thm i1.a.asm_A thm LocaleTest.IA_locale.i1.a.asm_A thm IA_locale.i1.a.asm_A | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 119 | thm i1.asm_A thm LocaleTest.i1.asm_A | 
| 15696 | 120 | |
| 17139 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 121 | ML {* check_thm "i1.a.asm_A" *}
 | 
| 15624 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 ballarin parents: 
15598diff
changeset | 122 | |
| 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 ballarin parents: 
15598diff
changeset | 123 | (* without prefix *) | 
| 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 ballarin parents: 
15598diff
changeset | 124 | |
| 27761 
b95e9ba0ca1d
Interpretation command (theory/proof context) no longer simplifies goal.
 ballarin parents: 
27718diff
changeset | 125 | interpretation IC ["W::i" "Z::i"] by intro_locales (* subsumed by i1: IC *) | 
| 19984 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 ballarin parents: 
19931diff
changeset | 126 | interpretation IC ["W::'a" "Z::i"] by unfold_locales auto | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 127 | (* subsumes i1: IA and i1: IC *) | 
| 15696 | 128 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 129 | print_interps IA (* output: <no prefix>, i1 *) | 
| 15596 | 130 | |
| 15837 | 131 | (* possible accesses *) | 
| 26645 
e114be97befe
Changed naming scheme for theorems generated by interpretations.
 ballarin parents: 
26343diff
changeset | 132 | thm asm_C thm a_b.asm_C thm LocaleTest.IC_locale.a_b.asm_C thm IC_locale.a_b.asm_C | 
| 15696 | 133 | |
| 17139 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 134 | ML {* check_thm "asm_C" *}
 | 
| 15596 | 135 | |
| 19931 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19783diff
changeset | 136 | interpretation i2: ID [X "Y::i" "Y = X"] | 
| 19984 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 ballarin parents: 
19931diff
changeset | 137 | by (simp add: eq_commute) unfold_locales | 
| 15837 | 138 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 139 | print_interps IA (* output: <no prefix>, i1 *) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 140 | print_interps ID (* output: i2 *) | 
| 15837 | 141 | |
| 142 | ||
| 19984 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 ballarin parents: 
19931diff
changeset | 143 | interpretation i3: ID [X "Y::i"] by simp unfold_locales | 
| 15598 
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
 ballarin parents: 
15596diff
changeset | 144 | |
| 17436 | 145 | (* duplicate: thm not added *) | 
| 15696 | 146 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 147 | (* thm i3.a.asm_A *) | 
| 15696 | 148 | |
| 15596 | 149 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 150 | print_interps IA (* output: <no prefix>, i1 *) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 151 | print_interps IB (* output: i1 *) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 152 | print_interps IC (* output: <no prefix, i1 *) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 153 | print_interps ID (* output: i2, i3 *) | 
| 15596 | 154 | |
| 155 | ||
| 27761 
b95e9ba0ca1d
Interpretation command (theory/proof context) no longer simplifies goal.
 ballarin parents: 
27718diff
changeset | 156 | interpretation i10: ID + ID a' b' d' [X "Y::i" _ u "v::i" _] by intro_locales | 
| 15696 | 157 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 158 | corollary (in ID) th_x: True .. | 
| 15696 | 159 | |
| 160 | (* possible accesses: for each registration *) | |
| 161 | ||
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 162 | thm i2.th_x thm i3.th_x | 
| 15696 | 163 | |
| 17139 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 164 | ML {* check_thm "i2.th_x"; check_thm "i3.th_x" *}
 | 
| 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 165 | |
| 26199 | 166 | lemma (in ID) th_y: "d == (a = b)" by (rule d_def) | 
| 15696 | 167 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 168 | thm i2.th_y thm i3.th_y | 
| 15696 | 169 | |
| 17139 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 170 | ML {* check_thm "i2.th_y"; check_thm "i3.th_y" *}
 | 
| 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 171 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 172 | lemmas (in ID) th_z = th_y | 
| 15696 | 173 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 174 | thm i2.th_z | 
| 15696 | 175 | |
| 17139 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 176 | ML {* check_thm "i2.th_z" *}
 | 
| 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 177 | |
| 15596 | 178 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 179 | subsection {* Interpretation in Proof Contexts *}
 | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 180 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 181 | locale IF = fixes f assumes asm_F: "f & f --> f" | 
| 16102 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
15837diff
changeset | 182 | |
| 25282 | 183 | consts default :: "'a" | 
| 184 | ||
| 15624 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 ballarin parents: 
15598diff
changeset | 185 | theorem True | 
| 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 ballarin parents: 
15598diff
changeset | 186 | proof - | 
| 25282 | 187 | fix alpha::i and beta | 
| 188 | have alpha_A: "IA(alpha)" by unfold_locales | |
| 27761 
b95e9ba0ca1d
Interpretation command (theory/proof context) no longer simplifies goal.
 ballarin parents: 
27718diff
changeset | 189 | interpret i5: IA [alpha] by intro_locales (* subsumed *) | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 190 | print_interps IA (* output: <no prefix>, i1 *) | 
| 19984 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 ballarin parents: 
19931diff
changeset | 191 | interpret i6: IC [alpha beta] by unfold_locales auto | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 192 | print_interps IA (* output: <no prefix>, i1 *) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 193 | print_interps IC (* output: <no prefix>, i1, i6 *) | 
| 25282 | 194 | interpret i11: IF ["default = default"] by (fast intro: IF.intro) | 
| 195 | thm i11.asm_F [where 'a = i] (* default has schematic type *) | |
| 15624 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 ballarin parents: 
15598diff
changeset | 196 | qed rule | 
| 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 ballarin parents: 
15598diff
changeset | 197 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 198 | theorem (in IA) True | 
| 15696 | 199 | proof - | 
| 19931 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19783diff
changeset | 200 | print_interps! IA | 
| 15696 | 201 | fix beta and gamma | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 202 | interpret i9: ID [a beta _] | 
| 19931 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19783diff
changeset | 203 | apply - apply assumption | 
| 19984 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 ballarin parents: 
19931diff
changeset | 204 | apply unfold_locales | 
| 19931 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19783diff
changeset | 205 | apply (rule refl) done | 
| 15696 | 206 | qed rule | 
| 15624 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 ballarin parents: 
15598diff
changeset | 207 | |
| 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 ballarin parents: 
15598diff
changeset | 208 | |
| 15696 | 209 | (* Definition involving free variable *) | 
| 210 | ||
| 211 | ML {* reset show_sorts *}
 | |
| 212 | ||
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 213 | locale IE = fixes e defines e_def: "e(x) == x & x" | 
| 15696 | 214 | notes e_def2 = e_def | 
| 215 | ||
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 216 | lemma (in IE) True thm e_def by fast | 
| 15696 | 217 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 218 | interpretation i7: IE ["%x. x"] by simp | 
| 15696 | 219 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 220 | thm i7.e_def2 (* has no premise *) | 
| 15696 | 221 | |
| 17139 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 222 | ML {* check_thm "i7.e_def2" *}
 | 
| 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 223 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 224 | locale IE' = fixes e defines e_def: "e == (%x. x & x)" | 
| 15696 | 225 | notes e_def2 = e_def | 
| 226 | ||
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 227 | interpretation i7': IE' ["(%x. x)"] by simp | 
| 15696 | 228 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 229 | thm i7'.e_def2 | 
| 15696 | 230 | |
| 17139 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 231 | ML {* check_thm "i7'.e_def2" *}
 | 
| 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 232 | |
| 15696 | 233 | (* Definition involving free variable in assm *) | 
| 234 | ||
| 27681 | 235 | locale IG = fixes g assumes asm_G: "g --> x" | 
| 16102 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
15837diff
changeset | 236 | notes asm_G2 = asm_G | 
| 15696 | 237 | |
| 27718 | 238 | interpretation i8: IG ["False"] by unfold_locales fast | 
| 15696 | 239 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 240 | thm i8.asm_G2 | 
| 15696 | 241 | |
| 17139 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 242 | ML {* check_thm "i8.asm_G2" *}
 | 
| 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 243 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 244 | text {* Locale without assumptions *}
 | 
| 15696 | 245 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 246 | locale IL1 = notes rev_conjI [intro] = conjI [THEN iffD1 [OF conj_commute]] | 
| 15696 | 247 | |
| 248 | lemma "[| P; Q |] ==> P & Q" | |
| 249 | proof - | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 250 |   interpret my: IL1 .          txt {* No chained fact required. *}
 | 
| 15696 | 251 |   assume Q and P               txt {* order reversed *}
 | 
| 252 |   then show "P & Q" ..         txt {* Applies @{thm my.rev_conjI}. *}
 | |
| 253 | qed | |
| 254 | ||
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 255 | locale IL11 = notes rev_conjI = conjI [THEN iffD1 [OF conj_commute]] | 
| 15696 | 256 | |
| 257 | ||
| 258 | subsection {* Simple locale with assumptions *}
 | |
| 259 | ||
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 260 | consts ibin :: "[i, i] => i" (infixl "#" 60) | 
| 15696 | 261 | |
| 262 | axioms i_assoc: "(x # y) # z = x # (y # z)" | |
| 263 | i_comm: "x # y = y # x" | |
| 264 | ||
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 265 | locale IL2 = | 
| 15696 | 266 | fixes OP (infixl "+" 60) | 
| 267 | assumes assoc: "(x + y) + z = x + (y + z)" | |
| 268 | and comm: "x + y = y + x" | |
| 269 | ||
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 270 | lemma (in IL2) lcomm: "x + (y + z) = y + (x + z)" | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 271 | proof - | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 272 | have "x + (y + z) = (x + y) + z" by (simp add: assoc) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 273 | also have "... = (y + x) + z" by (simp add: comm) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 274 | also have "... = y + (x + z)" by (simp add: assoc) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 275 | finally show ?thesis . | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 276 | qed | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 277 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 278 | lemmas (in IL2) AC = comm assoc lcomm | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 279 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 280 | lemma "(x::i) # y # z # w = y # x # w # z" | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 281 | proof - | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 282 | interpret my: IL2 ["op #"] by (rule IL2.intro [of "op #", OF i_assoc i_comm]) | 
| 17033 | 283 | show ?thesis by (simp only: my.OP.AC) (* or my.AC *) | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 284 | qed | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 285 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 286 | subsection {* Nested locale with assumptions *}
 | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 287 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 288 | locale IL3 = | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 289 | fixes OP (infixl "+" 60) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 290 | assumes assoc: "(x + y) + z = x + (y + z)" | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 291 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 292 | locale IL4 = IL3 + | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 293 | assumes comm: "x + y = y + x" | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 294 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 295 | lemma (in IL4) lcomm: "x + (y + z) = y + (x + z)" | 
| 15696 | 296 | proof - | 
| 297 | have "x + (y + z) = (x + y) + z" by (simp add: assoc) | |
| 298 | also have "... = (y + x) + z" by (simp add: comm) | |
| 299 | also have "... = y + (x + z)" by (simp add: assoc) | |
| 300 | finally show ?thesis . | |
| 301 | qed | |
| 302 | ||
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 303 | lemmas (in IL4) AC = comm assoc lcomm | 
| 15696 | 304 | |
| 305 | lemma "(x::i) # y # z # w = y # x # w # z" | |
| 306 | proof - | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 307 | interpret my: IL4 ["op #"] | 
| 19931 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19783diff
changeset | 308 | by (auto intro: IL4.intro IL3.intro IL4_axioms.intro i_assoc i_comm) | 
| 15696 | 309 | show ?thesis by (simp only: my.OP.AC) (* or simply AC *) | 
| 310 | qed | |
| 311 | ||
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 312 | text {* Locale with definition *}
 | 
| 15696 | 313 | |
| 314 | text {* This example is admittedly not very creative :-) *}
 | |
| 315 | ||
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 316 | locale IL5 = IL4 + var A + | 
| 15696 | 317 | defines A_def: "A == True" | 
| 318 | ||
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 319 | lemma (in IL5) lem: A | 
| 15696 | 320 | by (unfold A_def) rule | 
| 321 | ||
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 322 | lemma "IL5(op #) ==> True" | 
| 15696 | 323 | proof - | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 324 | assume "IL5(op #)" | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 325 | then interpret IL5 ["op #"] by (auto intro: IL5.axioms) | 
| 15696 | 326 | show ?thesis by (rule lem) (* lem instantiated to True *) | 
| 327 | qed | |
| 328 | ||
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 329 | text {* Interpretation in a context with target *}
 | 
| 15696 | 330 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 331 | lemma (in IL4) | 
| 15696 | 332 | fixes A (infixl "$" 60) | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 333 | assumes A: "IL4(A)" | 
| 15696 | 334 | shows "(x::i) $ y $ z $ w = y $ x $ w $ z" | 
| 335 | proof - | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 336 | from A interpret A: IL4 ["A"] by (auto intro: IL4.axioms) | 
| 15696 | 337 | show ?thesis by (simp only: A.OP.AC) | 
| 338 | qed | |
| 339 | ||
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 340 | |
| 16736 
1e792b32abef
Preparations for interpretation of locales in locales.
 ballarin parents: 
16620diff
changeset | 341 | section {* Interpretation in Locales *}
 | 
| 
1e792b32abef
Preparations for interpretation of locales in locales.
 ballarin parents: 
16620diff
changeset | 342 | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 343 | text {* Naming convention for global objects: prefixes R and r *}
 | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 344 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 345 | (* locale with many parameters --- | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 346 | interpretations generate alternating group A5 *) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 347 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 348 | locale RA5 = var A + var B + var C + var D + var E + | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 349 | assumes eq: "A <-> B <-> C <-> D <-> E" | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 350 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 351 | interpretation RA5 < RA5 _ _ D E C | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 352 | print_facts | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 353 | print_interps RA5 | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 354 | using A_B_C_D_E.eq apply (blast intro: RA5.intro) done | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 355 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 356 | interpretation RA5 < RA5 C _ E _ A | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 357 | print_facts | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 358 | print_interps RA5 | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 359 | using A_B_C_D_E.eq apply (blast intro: RA5.intro) done | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 360 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 361 | interpretation RA5 < RA5 B C A _ _ | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 362 | print_facts | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 363 | print_interps RA5 | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 364 | using A_B_C_D_E.eq apply (blast intro: RA5.intro) done | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 365 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 366 | interpretation RA5 < RA5 _ C D B _ . | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 367 | (* Any even permutation of parameters is subsumed by the above. *) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 368 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 369 | (* circle of three locales, forward direction *) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 370 | |
| 27681 | 371 | locale RA1 = var A + var B + assumes p: "A <-> B" | 
| 372 | locale RA2 = var A + var B + assumes q: "A & B | ~ A & ~ B" | |
| 373 | locale RA3 = var A + var B + assumes r: "(A --> B) & (B --> A)" | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 374 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 375 | interpretation RA1 < RA2 | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 376 | print_facts | 
| 27681 | 377 | using p apply unfold_locales apply fast done | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 378 | interpretation RA2 < RA3 | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 379 | print_facts | 
| 27681 | 380 | using q apply unfold_locales apply fast done | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 381 | interpretation RA3 < RA1 | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 382 | print_facts | 
| 27681 | 383 | using r apply unfold_locales apply fast done | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 384 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 385 | (* circle of three locales, backward direction *) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 386 | |
| 27681 | 387 | locale RB1 = var A + var B + assumes p: "A <-> B" | 
| 388 | locale RB2 = var A + var B + assumes q: "A & B | ~ A & ~ B" | |
| 389 | locale RB3 = var A + var B + assumes r: "(A --> B) & (B --> A)" | |
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 390 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 391 | interpretation RB1 < RB2 | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 392 | print_facts | 
| 27681 | 393 | using p apply unfold_locales apply fast done | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 394 | interpretation RB3 < RB1 | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 395 | print_facts | 
| 27681 | 396 | using r apply unfold_locales apply fast done | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 397 | interpretation RB2 < RB3 | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 398 | print_facts | 
| 27681 | 399 | using q apply unfold_locales apply fast done | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 400 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 401 | lemma (in RB1) True | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 402 | print_facts | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 403 | .. | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 404 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 405 | |
| 27718 | 406 | (* Group example *) | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 407 | |
| 27718 | 408 | locale Rsemi = var prod (infixl "**" 65) + | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 409 | assumes assoc: "(x ** y) ** z = x ** (y ** z)" | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 410 | |
| 27718 | 411 | locale Rlgrp = Rsemi + var one + var inv + | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 412 | assumes lone: "one ** x = x" | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 413 | and linv: "inv(x) ** x = one" | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 414 | |
| 27718 | 415 | lemma (in Rlgrp) lcancel: | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 416 | "x ** y = x ** z <-> y = z" | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 417 | proof | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 418 | assume "x ** y = x ** z" | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 419 | then have "inv(x) ** x ** y = inv(x) ** x ** z" by (simp add: assoc) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 420 | then show "y = z" by (simp add: lone linv) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 421 | qed simp | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 422 | |
| 27718 | 423 | locale Rrgrp = Rsemi + var one + var inv + | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 424 | assumes rone: "x ** one = x" | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 425 | and rinv: "x ** inv(x) = one" | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 426 | |
| 27718 | 427 | lemma (in Rrgrp) rcancel: | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 428 | "y ** x = z ** x <-> y = z" | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 429 | proof | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 430 | assume "y ** x = z ** x" | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 431 | then have "y ** (x ** inv(x)) = z ** (x ** inv(x))" | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 432 | by (simp add: assoc [symmetric]) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 433 | then show "y = z" by (simp add: rone rinv) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 434 | qed simp | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 435 | |
| 27718 | 436 | interpretation Rlgrp < Rrgrp | 
| 19984 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 ballarin parents: 
19931diff
changeset | 437 | proof unfold_locales | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 438 |     {
 | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 439 | fix x | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 440 | have "inv(x) ** x ** one = inv(x) ** x" by (simp add: linv lone) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 441 | then show "x ** one = x" by (simp add: assoc lcancel) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 442 | } | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 443 | note rone = this | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 444 |     {
 | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 445 | fix x | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 446 | have "inv(x) ** x ** inv(x) = inv(x) ** one" | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 447 | by (simp add: linv lone rone) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 448 | then show "x ** inv(x) = one" by (simp add: assoc lcancel) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 449 | } | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 450 | qed | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 451 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 452 | (* effect on printed locale *) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 453 | |
| 27718 | 454 | print_locale! Rlgrp | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 455 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 456 | (* use of derived theorem *) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 457 | |
| 27718 | 458 | lemma (in Rlgrp) | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 459 | "y ** x = z ** x <-> y = z" | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 460 | apply (rule rcancel) | 
| 27718 | 461 | print_interps Rrgrp thm lcancel rcancel | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 462 | done | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 463 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 464 | (* circular interpretation *) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 465 | |
| 27718 | 466 | interpretation Rrgrp < Rlgrp | 
| 19984 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 ballarin parents: 
19931diff
changeset | 467 | proof unfold_locales | 
| 17000 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 468 |     {
 | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 469 | fix x | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 470 | have "one ** (x ** inv(x)) = x ** inv(x)" by (simp add: rinv rone) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 471 | then show "one ** x = x" by (simp add: assoc [symmetric] rcancel) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 472 | } | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 473 | note lone = this | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 474 |     {
 | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 475 | fix x | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 476 | have "inv(x) ** (x ** inv(x)) = one ** inv(x)" | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 477 | by (simp add: rinv lone rone) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 478 | then show "inv(x) ** x = one" by (simp add: assoc [symmetric] rcancel) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 479 | } | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 480 | qed | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 481 | |
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 482 | (* effect on printed locale *) | 
| 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 ballarin parents: 
16736diff
changeset | 483 | |
| 27718 | 484 | print_locale! Rrgrp | 
| 485 | print_locale! Rlgrp | |
| 16736 
1e792b32abef
Preparations for interpretation of locales in locales.
 ballarin parents: 
16620diff
changeset | 486 | |
| 17033 | 487 | subsection {* Interaction of Interpretation in Theories and Locales:
 | 
| 17096 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 488 | in Locale, then in Theory *} | 
| 17033 | 489 | |
| 490 | consts | |
| 491 | rone :: i | |
| 492 | rinv :: "i => i" | |
| 493 | ||
| 494 | axioms | |
| 495 | r_one : "rone # x = x" | |
| 496 | r_inv : "rinv(x) # x = rone" | |
| 497 | ||
| 498 | interpretation Rbool: Rlgrp ["op #" "rone" "rinv"] | |
| 27681 | 499 | proof | 
| 17033 | 500 | fix x y z | 
| 501 |   {
 | |
| 502 | show "(x # y) # z = x # (y # z)" by (rule i_assoc) | |
| 503 | next | |
| 504 | show "rone # x = x" by (rule r_one) | |
| 505 | next | |
| 506 | show "rinv(x) # x = rone" by (rule r_inv) | |
| 507 | } | |
| 508 | qed | |
| 509 | ||
| 510 | (* derived elements *) | |
| 511 | ||
| 512 | print_interps Rrgrp | |
| 513 | print_interps Rlgrp | |
| 514 | ||
| 515 | lemma "y # x = z # x <-> y = z" by (rule Rbool.rcancel) | |
| 516 | ||
| 517 | (* adding lemma to derived element *) | |
| 518 | ||
| 519 | lemma (in Rrgrp) new_cancel: | |
| 520 | "b ** a = c ** a <-> b = c" | |
| 521 | by (rule rcancel) | |
| 522 | ||
| 523 | thm Rbool.new_cancel (* additional prems discharged!! *) | |
| 524 | ||
| 17139 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 525 | ML {* check_thm "Rbool.new_cancel" *}
 | 
| 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 526 | |
| 17033 | 527 | lemma "b # a = c # a <-> b = c" by (rule Rbool.new_cancel) | 
| 528 | ||
| 529 | ||
| 530 | subsection {* Interaction of Interpretation in Theories and Locales:
 | |
| 17096 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 531 | in Theory, then in Locale *} | 
| 17033 | 532 | |
| 533 | (* Another copy of the group example *) | |
| 534 | ||
| 535 | locale Rqsemi = var prod (infixl "**" 65) + | |
| 536 | assumes assoc: "(x ** y) ** z = x ** (y ** z)" | |
| 537 | ||
| 538 | locale Rqlgrp = Rqsemi + var one + var inv + | |
| 539 | assumes lone: "one ** x = x" | |
| 540 | and linv: "inv(x) ** x = one" | |
| 541 | ||
| 542 | lemma (in Rqlgrp) lcancel: | |
| 543 | "x ** y = x ** z <-> y = z" | |
| 544 | proof | |
| 545 | assume "x ** y = x ** z" | |
| 546 | then have "inv(x) ** x ** y = inv(x) ** x ** z" by (simp add: assoc) | |
| 547 | then show "y = z" by (simp add: lone linv) | |
| 548 | qed simp | |
| 549 | ||
| 550 | locale Rqrgrp = Rqsemi + var one + var inv + | |
| 551 | assumes rone: "x ** one = x" | |
| 552 | and rinv: "x ** inv(x) = one" | |
| 553 | ||
| 554 | lemma (in Rqrgrp) rcancel: | |
| 555 | "y ** x = z ** x <-> y = z" | |
| 556 | proof | |
| 557 | assume "y ** x = z ** x" | |
| 558 | then have "y ** (x ** inv(x)) = z ** (x ** inv(x))" | |
| 559 | by (simp add: assoc [symmetric]) | |
| 560 | then show "y = z" by (simp add: rone rinv) | |
| 561 | qed simp | |
| 562 | ||
| 27718 | 563 | interpretation Rqrgrp < Rrgrp | 
| 19984 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 ballarin parents: 
19931diff
changeset | 564 | apply unfold_locales | 
| 19931 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19783diff
changeset | 565 | apply (rule assoc) | 
| 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19783diff
changeset | 566 | apply (rule rone) | 
| 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19783diff
changeset | 567 | apply (rule rinv) | 
| 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19783diff
changeset | 568 | done | 
| 17139 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 569 | |
| 17033 | 570 | interpretation R2: Rqlgrp ["op #" "rone" "rinv"] | 
| 19984 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 ballarin parents: 
19931diff
changeset | 571 | apply unfold_locales (* FIXME: unfold_locales is too eager and shouldn't | 
| 19931 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19783diff
changeset | 572 | solve this. *) | 
| 27681 | 573 | apply (rule i_assoc) | 
| 574 | apply (rule r_one) | |
| 575 | apply (rule r_inv) | |
| 19931 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19783diff
changeset | 576 | done | 
| 17033 | 577 | |
| 578 | print_interps Rqsemi | |
| 579 | print_interps Rqlgrp | |
| 27718 | 580 | print_interps Rlgrp (* no interpretations yet *) | 
| 17139 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 581 | |
| 17033 | 582 | |
| 583 | interpretation Rqlgrp < Rqrgrp | |
| 19984 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 ballarin parents: 
19931diff
changeset | 584 | proof unfold_locales | 
| 17033 | 585 |     {
 | 
| 586 | fix x | |
| 587 | have "inv(x) ** x ** one = inv(x) ** x" by (simp add: linv lone) | |
| 588 | then show "x ** one = x" by (simp add: assoc lcancel) | |
| 589 | } | |
| 590 | note rone = this | |
| 591 |     {
 | |
| 592 | fix x | |
| 593 | have "inv(x) ** x ** inv(x) = inv(x) ** one" | |
| 594 | by (simp add: linv lone rone) | |
| 595 | then show "x ** inv(x) = one" by (simp add: assoc lcancel) | |
| 596 | } | |
| 597 | qed | |
| 598 | ||
| 17139 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 599 | print_interps! Rqrgrp | 
| 27718 | 600 | print_interps! Rsemi (* witness must not have meta hyps *) | 
| 601 | print_interps! Rrgrp (* witness must not have meta hyps *) | |
| 602 | print_interps! Rlgrp (* witness must not have meta hyps *) | |
| 17139 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 603 | thm R2.rcancel | 
| 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 604 | thm R2.lcancel | 
| 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 605 | |
| 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 606 | ML {* check_thm "R2.rcancel"; check_thm "R2.lcancel" *}
 | 
| 
165c97f9bb63
Printing of interpretations: option to show witness theorems;
 ballarin parents: 
17096diff
changeset | 607 | |
| 17096 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 608 | |
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 609 | subsection {* Generation of Witness Theorems for Transitive Interpretations *}
 | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 610 | |
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 611 | locale Rtriv = var x + | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 612 | assumes x: "x = x" | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 613 | |
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 614 | locale Rtriv2 = var x + var y + | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 615 | assumes x: "x = x" and y: "y = y" | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 616 | |
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 617 | interpretation Rtriv2 < Rtriv x | 
| 19984 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 ballarin parents: 
19931diff
changeset | 618 | apply unfold_locales | 
| 17096 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 619 | apply (rule x) | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 620 | done | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 621 | |
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 622 | interpretation Rtriv2 < Rtriv y | 
| 19984 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 ballarin parents: 
19931diff
changeset | 623 | apply unfold_locales | 
| 17096 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 624 | apply (rule y) | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 625 | done | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 626 | |
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 627 | print_locale Rtriv2 | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 628 | |
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 629 | locale Rtriv3 = var x + var y + var z + | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 630 | assumes x: "x = x" and y: "y = y" and z: "z = z" | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 631 | |
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 632 | interpretation Rtriv3 < Rtriv2 x y | 
| 19984 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 ballarin parents: 
19931diff
changeset | 633 | apply unfold_locales | 
| 17096 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 634 | apply (rule x) | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 635 | apply (rule y) | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 636 | done | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 637 | |
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 638 | print_locale Rtriv3 | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 639 | |
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 640 | interpretation Rtriv3 < Rtriv2 x z | 
| 19984 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 ballarin parents: 
19931diff
changeset | 641 | apply unfold_locales | 
| 17096 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 642 | apply (rule x_y_z.x) | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 643 | apply (rule z) | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 644 | done | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 645 | |
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 646 | ML {* set show_types *}
 | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 647 | |
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 648 | print_locale Rtriv3 | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 649 | |
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 650 | |
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 651 | subsection {* Normalisation Replaces Assumed Element by Derived Element *}
 | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 652 | |
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 653 | typedecl ('a, 'b) pair
 | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 654 | arities pair :: ("term", "term") "term"
 | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 655 | |
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 656 | consts | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 657 |   pair :: "['a, 'b] => ('a, 'b) pair"
 | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 658 |   fst :: "('a, 'b) pair => 'a"
 | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 659 |   snd :: "('a, 'b) pair => 'b"
 | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 660 | |
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 661 | axioms | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 662 | fst [simp]: "fst(pair(x, y)) = x" | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 663 | snd [simp]: "snd(pair(x, y)) = y" | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 664 | |
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 665 | locale Rpair = var prod (infixl "**" 65) + var prodP (infixl "***" 65) + | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 666 | defines P_def: "x *** y == pair(fst(x) ** fst(y), snd(x) ** snd(y))" | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 667 | |
| 27718 | 668 | locale Rpair_semi = Rpair + Rsemi | 
| 17096 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 669 | |
| 27718 | 670 | interpretation Rpair_semi < Rsemi prodP (infixl "***" 65) | 
| 671 | proof unfold_locales | |
| 17096 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 672 | fix x y z | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 673 | show "(x *** y) *** z = x *** (y *** z)" | 
| 19931 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19783diff
changeset | 674 | apply (simp only: P_def) apply (simp add: assoc) (* FIXME: unfold P_def fails *) | 
| 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19783diff
changeset | 675 | done | 
| 17096 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 676 | qed | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 677 | |
| 27718 | 678 | locale Rsemi_rev = Rsemi + var rprod (infixl "++" 65) + | 
| 17096 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 679 | defines r_def: "x ++ y == y ** x" | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 680 | |
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 681 | lemma (in Rsemi_rev) r_assoc: | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 682 | "(x ++ y) ++ z = x ++ (y ++ z)" | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 683 | by (simp add: r_def assoc) | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 684 | |
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 685 | lemma (in Rpair_semi) | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 686 | includes Rsemi_rev prodP (infixl "***" 65) rprodP (infixl "+++" 65) | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 687 | constrains prod :: "['a, 'a] => 'a" | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 688 |     and rprodP :: "[('a, 'a) pair, ('a, 'a) pair] => ('a, 'a) pair"
 | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 689 | shows "(x +++ y) +++ z = x +++ (y +++ z)" | 
| 
8327b71282ce
Improved generation of witnesses in interpretation.
 ballarin parents: 
17033diff
changeset | 690 | apply (rule r_assoc) done | 
| 17033 | 691 | |
| 20469 | 692 | |
| 693 | subsection {* Import of Locales with Predicates as Interpretation *}
 | |
| 694 | ||
| 695 | locale Ra = | |
| 696 | assumes Ra: "True" | |
| 697 | ||
| 698 | locale Rb = Ra + | |
| 699 | assumes Rb: "True" | |
| 700 | ||
| 701 | locale Rc = Rb + | |
| 702 | assumes Rc: "True" | |
| 703 | ||
| 704 | print_locale! Rc | |
| 705 | ||
| 22659 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 706 | |
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 707 | section {* Interpretation of Defined Concepts *}
 | 
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 708 | |
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 709 | text {* Naming convention for global objects: prefixes D and d *}
 | 
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 710 | |
| 23919 
af871d13e320
interpretation: equations are propositions not pairs of terms;
 ballarin parents: 
22931diff
changeset | 711 | |
| 
af871d13e320
interpretation: equations are propositions not pairs of terms;
 ballarin parents: 
22931diff
changeset | 712 | subsection {* Simple examples *}
 | 
| 
af871d13e320
interpretation: equations are propositions not pairs of terms;
 ballarin parents: 
22931diff
changeset | 713 | |
| 22659 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 714 | locale Da = fixes a :: o | 
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 715 | assumes true: a | 
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 716 | |
| 22757 | 717 | text {* In the following examples, @{term "~ a"} is the defined concept. *}
 | 
| 718 | ||
| 22659 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 719 | lemma (in Da) not_false: "~ a <-> False" | 
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 720 | apply simp apply (rule true) done | 
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 721 | |
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 722 | interpretation D1: Da ["True"] | 
| 23919 
af871d13e320
interpretation: equations are propositions not pairs of terms;
 ballarin parents: 
22931diff
changeset | 723 | where "~ True == False" | 
| 22659 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 724 | apply - | 
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 725 | apply unfold_locales [1] apply rule | 
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 726 | by simp | 
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 727 | |
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 728 | thm D1.not_false | 
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 729 | lemma "False <-> False" apply (rule D1.not_false) done | 
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 730 | |
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 731 | interpretation D2: Da ["x | ~ x"] | 
| 23919 
af871d13e320
interpretation: equations are propositions not pairs of terms;
 ballarin parents: 
22931diff
changeset | 732 | where "~ (x | ~ x) <-> ~ x & x" | 
| 22659 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 733 | apply - | 
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 734 | apply unfold_locales [1] apply fast | 
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 735 | by simp | 
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 736 | |
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 737 | thm D2.not_false | 
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 738 | lemma "~ x & x <-> False" apply (rule D2.not_false) done | 
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 739 | |
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 740 | print_interps! Da | 
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 741 | |
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 742 | (* Subscriptions of interpretations *) | 
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 743 | |
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 744 | lemma (in Da) not_false2: "~a <-> False" | 
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 745 | apply simp apply (rule true) done | 
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 746 | |
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 747 | thm D1.not_false2 D2.not_false2 | 
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 748 | lemma "False <-> False" apply (rule D1.not_false2) done | 
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 749 | lemma "~x & x <-> False" apply (rule D2.not_false2) done | 
| 
f792579b6e59
Experimental code for the interpretation of definitions.
 ballarin parents: 
20469diff
changeset | 750 | |
| 22757 | 751 | (* Unfolding in attributes *) | 
| 752 | ||
| 753 | locale Db = Da + | |
| 754 | fixes b :: o | |
| 755 | assumes a_iff_b: "~a <-> b" | |
| 756 | ||
| 757 | lemmas (in Db) not_false_b = not_false [unfolded a_iff_b] | |
| 758 | ||
| 759 | interpretation D2: Db ["x | ~ x" "~ (x <-> x)"] | |
| 760 | apply unfold_locales apply fast done | |
| 761 | ||
| 762 | thm D2.not_false_b | |
| 763 | lemma "~(x <-> x) <-> False" apply (rule D2.not_false_b) done | |
| 764 | ||
| 765 | (* Subscription and attributes *) | |
| 766 | ||
| 767 | lemmas (in Db) not_false_b2 = not_false [unfolded a_iff_b] | |
| 768 | ||
| 769 | thm D2.not_false_b2 | |
| 770 | lemma "~(x <-> x) <-> False" apply (rule D2.not_false_b2) done | |
| 771 | ||
| 15596 | 772 | end |