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